From 3367bbe28257b4bb2aadbbe15247ad5273c0c720 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Tue, 19 Nov 2024 07:55:06 +0000 Subject: [PATCH 001/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6125 --- 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..bf4ddd0bf6edc 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "23b4216d1e3a2064058bb415c11e5f6df0417c99", + "rev": "07d6e1c0dc621c47aa7cc9e8b9d8ba814d26ece5", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6125", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..b9475392717a6 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-6125" 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..9bd884f4630e7 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-18 +leanprover/lean4-pr-releases:pr-release-6125 From c7a792c828bf9d2728ef09ad0264a2861d5da4f4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 04:59:12 +0000 Subject: [PATCH 002/250] chore: adaptations for nightly-2024-11-20 (#19314) Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Johan Commelin Co-authored-by: Kyle Miller --- Mathlib/Analysis/Calculus/FDeriv/Mul.lean | 4 ++-- .../Analysis/FunctionalSpaces/SobolevInequality.lean | 10 ++++++++-- 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/Sym.lean | 2 +- Mathlib/Data/List/TFAE.lean | 2 +- Mathlib/Data/Vector/Mem.lean | 4 +--- Mathlib/Lean/Meta/KAbstractPositions.lean | 2 +- .../ClassNumber/AdmissibleAbsoluteValue.lean | 2 +- Mathlib/Tactic/CC/Addition.lean | 2 +- Mathlib/Tactic/FBinop.lean | 4 ++-- Mathlib/Tactic/Linarith/Frontend.lean | 2 +- Mathlib/Tactic/Linter/PPRoundtrip.lean | 2 +- Mathlib/Tactic/MoveAdd.lean | 3 ++- Mathlib/Tactic/ToAdditive/Frontend.lean | 2 +- Mathlib/Tactic/Variable.lean | 4 ++-- Mathlib/Testing/Plausible/Functions.lean | 2 +- lake-manifest.json | 6 +++--- lean-toolchain | 2 +- 21 files changed, 40 insertions(+), 35 deletions(-) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean index e8073446374b3..3bc95aa8d8223 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Mul.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Mul.lean @@ -678,7 +678,7 @@ theorem HasFDerivAt.list_prod' {l : List ι} {x : E} smulRight (f' l[i]) ((l.drop (.succ i)).map (f · x)).prod) x := by simp only [← List.finRange_map_get l, List.map_map] refine .congr_fderiv (hasFDerivAt_list_prod_finRange'.comp x - (hasFDerivAt_pi.mpr fun i ↦ h l[i] (List.getElem_mem i.isLt))) ?_ + (hasFDerivAt_pi.mpr fun i ↦ h l[i] (l.getElem_mem _))) ?_ ext m simp [← List.map_map] @@ -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 _))) ?_ ext m simp [← List.map_map] diff --git a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean index 12f3131cc2e8b..e18e2e98d3690 100644 --- a/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean +++ b/Mathlib/Analysis/FunctionalSpaces/SobolevInequality.lean @@ -680,7 +680,10 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] have hp' : p'⁻¹ = p⁻¹ - (finrank ℝ E : ℝ)⁻¹ := by rw [inv_inv, NNReal.coe_sub] · simp - · gcongr + · #adaptation_note + /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ + exact inv_anti₀ (by positivity) h2p.le have : (q : ℝ≥0∞) ≤ p' := by have H : (p' : ℝ)⁻¹ ≤ (↑q)⁻¹ := trans hp' hpq norm_cast at H ⊢ @@ -688,7 +691,10 @@ theorem eLpNorm_le_eLpNorm_fderiv_of_le [FiniteDimensional ℝ F] · dsimp have : 0 < p⁻¹ - (finrank ℝ E : ℝ≥0)⁻¹ := by simp only [tsub_pos_iff_lt] - gcongr + #adaptation_note + /-- This should just be `gcongr`, but this is not working as of nightly-2024-11-20. + Possibly related to #19262 (since this proof fails at `with_reducible_and_instances`). -/ + exact inv_strictAnti₀ (by positivity) h2p positivity · positivity set t := (μ s).toNNReal ^ (1 / q - 1 / p' : ℝ) diff --git a/Mathlib/Combinatorics/Enumerative/Composition.lean b/Mathlib/Combinatorics/Enumerative/Composition.lean index 986e40eed5b1b..e1bf3970d658a 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 _ _ _ + get_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 (get_mem (blocks c) _) @[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..d4354b5039a5b 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.get_mem l _) 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.get_mem l _) 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..34485c3ee467f 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 (get_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) (get_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 get_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 get_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) (get_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..b355f9e72ed52 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, get_mem _ _⟩ 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/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/Data/Vector/Mem.lean b/Mathlib/Data/Vector/Mem.lean index 3eb0fe6678d82..dc09c5bb58415 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.get_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/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/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean b/Mathlib/NumberTheory/ClassNumber/AdmissibleAbsoluteValue.lean index bdc2a13623ee7..0d34bda52b147 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.get_mem _ ⟨i, 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), 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/FBinop.lean b/Mathlib/Tactic/FBinop.lean index 3eb83fec8d8d6..8e58cb0849763 100644 --- a/Mathlib/Tactic/FBinop.lean +++ b/Mathlib/Tactic/FBinop.lean @@ -185,12 +185,12 @@ 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 <| withTermInfoContext' .anonymous ref do let lhs ← toExprCore lhs let mut rhs ← toExprCore rhs mkBinOp f lhs rhs | .macroExpansion macroName stx stx' nested => - withRef stx <| withInfoContext' stx (mkInfo := mkTermInfo macroName stx) do + withRef stx <| withTermInfoContext' macroName stx do withMacroExpansion stx stx' do toExprCore nested diff --git a/Mathlib/Tactic/Linarith/Frontend.lean b/Mathlib/Tactic/Linarith/Frontend.lean index 04b1e07d1ac97..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 diff --git a/Mathlib/Tactic/Linter/PPRoundtrip.lean b/Mathlib/Tactic/Linter/PPRoundtrip.lean index 9f8e258ecd46f..df238e0927ee3 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.eraseIdxIfInBounds 0 let s := (split.map .trimLeft).filter (· != "") (" ".intercalate (s.filter (!·.isEmpty)), preWS) 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 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/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/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 diff --git a/lake-manifest.json b/lake-manifest.json index 0df8444539094..4b7b5e02fb6e4 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": "0c995c64c9c8e4186e24f85d5d61bc404b378ba6", "name": "aesop", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9762792852464bf9591f6cccfe44f14b3ef54b25", + "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", 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 74e5aac4afb1d9247477d5e46add53480bc5809d Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Thu, 21 Nov 2024 05:30:11 +0000 Subject: [PATCH 003/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6139 --- 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 4b7b5e02fb6e4..fcc43c87eea7e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", + "rev": "8f31c34681b29b45deb5de421a601b3f39899432", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6139", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..637aa01e5f7e6 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-6139" 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 118d9e578af27..eb76ef6880344 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-20 +leanprover/lean4-pr-releases:pr-release-6139 From 9b780995ae4db58a6261a4d5a3ac95029969596f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 17:02:08 +1100 Subject: [PATCH 004/250] fixes for leanprover/lean4#6139 --- .../Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean index 9858e94c23f14..17812e9c56178 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean @@ -60,7 +60,7 @@ nonnegative. def checkSuccess : SimplexAlgorithmM matType Bool := do let lastIdx := (← get).free.size - 1 return (← get).mat[(0, lastIdx)]! > 0 && - (← Nat.allM (← get).basic.size (fun i => do return (← get).mat[(i, lastIdx)]! >= 0)) + (← (← get).basic.size.allM (fun i _ => do return (← get).mat[(i, lastIdx)]! ≥ 0)) /-- Chooses an entering variable: among the variables with a positive coefficient in the objective From 9c25cf14f11043d652fbbba93c8dbe09695a3dd3 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Thu, 21 Nov 2024 07:33:15 +0000 Subject: [PATCH 005/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6112 --- 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 4b7b5e02fb6e4..aafb1ceadefe5 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", + "rev": "fcc096fdce0fedca76f50f8e9d6add4eb5034962", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6112", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..f87f9bd15760d 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-6112" 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 118d9e578af27..625004eda4402 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-20 +leanprover/lean4-pr-releases:pr-release-6112 From c2fcf775f950f80d141260cc057afd4ffd81ae72 Mon Sep 17 00:00:00 2001 From: Jon Eugster Date: Thu, 21 Nov 2024 10:01:54 +0000 Subject: [PATCH 006/250] chore: rename `Nat.prime_def_lt''` to `Nat.prime_def` (#19255) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Rename the statement `Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p` to `Nat.prime_def`. Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Algebra/CharP/Defs.lean | 2 +- Mathlib/Data/Nat/Prime/Defs.lean | 11 ++++++++--- Mathlib/GroupTheory/SpecificGroups/Cyclic.lean | 2 +- Mathlib/NumberTheory/Divisors.lean | 2 +- Mathlib/NumberTheory/Fermat.lean | 2 +- scripts/no_lints_prime_decls.txt | 1 - 6 files changed, 12 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/CharP/Defs.lean b/Mathlib/Algebra/CharP/Defs.lean index eb523f7e90fac..093c600a4a310 100644 --- a/Mathlib/Algebra/CharP/Defs.lean +++ b/Mathlib/Algebra/CharP/Defs.lean @@ -232,7 +232,7 @@ section NoZeroDivisors variable [NoZeroDivisors R] lemma char_is_prime_of_two_le (p : ℕ) [CharP R p] (hp : 2 ≤ p) : Nat.Prime p := - suffices ∀ (d) (_ : d ∣ p), d = 1 ∨ d = p from Nat.prime_def_lt''.mpr ⟨hp, this⟩ + suffices ∀ (d) (_ : d ∣ p), d = 1 ∨ d = p from Nat.prime_def.mpr ⟨hp, this⟩ fun (d : ℕ) (hdvd : ∃ e, p = d * e) => let ⟨e, hmul⟩ := hdvd have : (p : R) = 0 := (cast_eq_zero_iff R p p).mpr (dvd_refl p) diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index e6a3e5bd1a467..30b5a96a46888 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -32,7 +32,8 @@ namespace Nat variable {n : ℕ} /-- `Nat.Prime p` means that `p` is a prime number, that is, a natural number - at least 2 whose only divisors are `p` and `1`. -/ + at least 2 whose only divisors are `p` and `1`. + The theorem `Nat.prime_def` witnesses this description of a prime number. -/ @[pp_nodot] def Prime (p : ℕ) := Irreducible p @@ -77,7 +78,8 @@ theorem Prime.eq_one_or_self_of_dvd {p : ℕ} (pp : p.Prime) (m : ℕ) (hm : m rintro rfl rw [hn, mul_one] -theorem prime_def_lt'' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p := by +@[inherit_doc Nat.Prime] +theorem prime_def {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m = 1 ∨ m = p := by refine ⟨fun h => ⟨h.two_le, h.eq_one_or_self_of_dvd⟩, fun h => ?_⟩ have h1 := Nat.one_lt_two.trans_le h.1 refine ⟨mt Nat.isUnit_iff.mp h1.ne', fun a b hab => ?_⟩ @@ -88,8 +90,11 @@ theorem prime_def_lt'' {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m, m ∣ p → m · rw [hab] exact dvd_mul_right _ _ +@[deprecated (since := "2024-11-19")] +alias prime_def_lt'' := prime_def + theorem prime_def_lt {p : ℕ} : Prime p ↔ 2 ≤ p ∧ ∀ m < p, m ∣ p → m = 1 := - prime_def_lt''.trans <| + prime_def.trans <| and_congr_right fun p2 => forall_congr' fun _ => ⟨fun h l d => (h d).resolve_right (ne_of_lt l), fun h d => diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index ca2a56cc761c7..11e9dc5991b44 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -551,7 +551,7 @@ instance (priority := 100) isCyclic : IsCyclic α := by theorem prime_card [Finite α] : (Nat.card α).Prime := by have h0 : 0 < Nat.card α := Nat.card_pos obtain ⟨g, hg⟩ := IsCyclic.exists_generator (α := α) - rw [Nat.prime_def_lt''] + rw [Nat.prime_def] refine ⟨Finite.one_lt_card_iff_nontrivial.2 inferInstance, fun n hn => ?_⟩ refine (IsSimpleOrder.eq_bot_or_eq_top (Subgroup.zpowers (g ^ n))).symm.imp ?_ ?_ · intro h diff --git a/Mathlib/NumberTheory/Divisors.lean b/Mathlib/NumberTheory/Divisors.lean index e2bfeafb8e217..5d34808cc4719 100644 --- a/Mathlib/NumberTheory/Divisors.lean +++ b/Mathlib/NumberTheory/Divisors.lean @@ -391,7 +391,7 @@ theorem Prime.prod_divisors {α : Type*} [CommMonoid α] {p : ℕ} {f : ℕ → theorem properDivisors_eq_singleton_one_iff_prime : n.properDivisors = {1} ↔ n.Prime := by refine ⟨?_, ?_⟩ · intro h - refine Nat.prime_def_lt''.mpr ⟨?_, fun m hdvd => ?_⟩ + refine Nat.prime_def.mpr ⟨?_, fun m hdvd => ?_⟩ · match n with | 0 => contradiction | 1 => contradiction diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index 18c9ed47758f2..bdb8dd1117803 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -196,7 +196,7 @@ theorem prime_of_pow_sub_one_prime {a n : ℕ} (hn1 : n ≠ 1) (hP : (a ^ n - 1) rw [one_pow, hP.dvd_iff_eq (mt (Nat.sub_eq_iff_eq_add ha1.le).mp hn1), eq_comm] at h exact (pow_eq_self_iff ha1).mp (Nat.sub_one_cancel ha0 (pow_pos ha0 n) h).symm subst ha2 - refine ⟨rfl, Nat.prime_def_lt''.mpr ⟨(two_le_iff n).mpr ⟨hn0, hn1⟩, fun d hdn ↦ ?_⟩⟩ + refine ⟨rfl, Nat.prime_def.mpr ⟨(two_le_iff n).mpr ⟨hn0, hn1⟩, fun d hdn ↦ ?_⟩⟩ have hinj : ∀ x y, 2 ^ x - 1 = 2 ^ y - 1 → x = y := fun x y h ↦ Nat.pow_right_injective le_rfl (sub_one_cancel (pow_pos ha0 x) (pow_pos ha0 y) h) let h := nat_sub_dvd_pow_sub_pow (2 ^ d) 1 (n / d) diff --git a/scripts/no_lints_prime_decls.txt b/scripts/no_lints_prime_decls.txt index e04b49b8dc0f0..fd905374bfebb 100644 --- a/scripts/no_lints_prime_decls.txt +++ b/scripts/no_lints_prime_decls.txt @@ -3315,7 +3315,6 @@ Nat.Partrec.rfind' Nat.pow_lt_ascFactorial' Nat.pow_sub_lt_descFactorial' Nat.prime_def_lt' -Nat.prime_def_lt'' Nat.Prime.eq_two_or_odd' Nat.primeFactorsList_chain' Nat.Prime.not_prime_pow' From 840e02ce2768e06de7ced0a624444746590e9d99 Mon Sep 17 00:00:00 2001 From: Pietro Monticone <38562595+pitmonticone@users.noreply.github.com> Date: Thu, 21 Nov 2024 10:33:44 +0000 Subject: [PATCH 007/250] feat(Data/Int/Interval): add `Icc_eq_pair` (#19308) Upstreamed from the [Carleson](https://github.com/fpvandoorn/carleson) project. Co-authored-by: James Sundstrom --- Mathlib/Data/Int/Interval.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Int/Interval.lean b/Mathlib/Data/Int/Interval.lean index a93f1ff49aa84..cd94ea340d1c5 100644 --- a/Mathlib/Data/Int/Interval.lean +++ b/Mathlib/Data/Int/Interval.lean @@ -129,6 +129,11 @@ theorem card_Ioc_of_le (h : a ≤ b) : (#(Ioc a b) : ℤ) = b - a := by theorem card_Ioo_of_lt (h : a < b) : (#(Ioo a b) : ℤ) = b - a - 1 := by rw [card_Ioo, sub_sub, toNat_sub_of_le h] +theorem Icc_eq_pair : Finset.Icc a (a + 1) = {a, a + 1} := by + ext + simp + omega + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it theorem card_fintype_Icc : Fintype.card (Set.Icc a b) = (b + 1 - a).toNat := by rw [← card_Icc, Fintype.card_ofFinset] From 77813d787779e3a6c179c575bed1d588f6b9c0d3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 21 Nov 2024 11:42:12 +0000 Subject: [PATCH 008/250] chore: adaptations for nightly-2024-11-21 (#19331) Co-authored-by: Kyle Miller Co-authored-by: JovanGerb Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Johan Commelin --- Mathlib/Tactic/WLOG.lean | 3 ++- lake-manifest.json | 4 ++-- lean-toolchain | 2 +- 3 files changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/WLOG.lean b/Mathlib/Tactic/WLOG.lean index 461949d04ab7b..cad9a19bf5bfd 100644 --- a/Mathlib/Tactic/WLOG.lean +++ b/Mathlib/Tactic/WLOG.lean @@ -77,7 +77,8 @@ def _root_.Lean.MVarId.wlog (goal : MVarId) (h : Option Name) (P : Expr) let (revertedFVars, HType) ← liftMkBindingM fun ctx => (do let f ← collectForwardDeps lctx fvars let revertedFVars := filterOutImplementationDetails lctx (f.map Expr.fvarId!) - let HType ← withFreshCache do mkAuxMVarType lctx (revertedFVars.map Expr.fvar) .natural HSuffix + let HType ← withFreshCache do + mkAuxMVarType lctx (revertedFVars.map Expr.fvar) .natural HSuffix (usedLetOnly := true) return (revertedFVars, HType)) { preserveOrder := false, mainModule := ctx.mainModule } /- Set up the goal which will suppose `h`; this begins as a goal with type H (hence HExpr), and h diff --git a/lake-manifest.json b/lake-manifest.json index 4b7b5e02fb6e4..4820e7097b0d2 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "9b9f4d0406d00baaae62d1a717b5aa854a2ae51d", + "rev": "b3935d53ce8dad5665af7ac41f06f0f6de4d942f", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6128", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lean-toolchain b/lean-toolchain index 118d9e578af27..f589f3bf8fa9a 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-20 +leanprover/lean4:nightly-2024-11-21 From bded1a9eb3fc57b04f6ff80b7449c55ed2a1baa4 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Thu, 21 Nov 2024 12:36:28 +0000 Subject: [PATCH 009/250] feat(RingTheory/DividedPowers/Basic): definition of divided powers (#15657) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Define the notion of a divided power structure on an ideal of a ring. Prove basic lemmas: * the variant of the multinomial theorem in this context * a formula for products of divided powers Co-authored with: [María Inés de Frutos-Fernández](https://github.com/mariainesdff) Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib.lean | 1 + Mathlib/RingTheory/DividedPowers/Basic.lean | 418 ++++++++++++++++++++ Mathlib/RingTheory/Ideal/Maps.lean | 11 + docs/references.bib | 47 +++ 4 files changed, 477 insertions(+) create mode 100644 Mathlib/RingTheory/DividedPowers/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 139b1e81351a5..e6b62736993a2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4197,6 +4197,7 @@ import Mathlib.RingTheory.Derivation.ToSquareZero import Mathlib.RingTheory.DiscreteValuationRing.Basic import Mathlib.RingTheory.DiscreteValuationRing.TFAE import Mathlib.RingTheory.Discriminant +import Mathlib.RingTheory.DividedPowers.Basic import Mathlib.RingTheory.DualNumber import Mathlib.RingTheory.EisensteinCriterion import Mathlib.RingTheory.EssentialFiniteness diff --git a/Mathlib/RingTheory/DividedPowers/Basic.lean b/Mathlib/RingTheory/DividedPowers/Basic.lean new file mode 100644 index 0000000000000..b07b6b0ac8aea --- /dev/null +++ b/Mathlib/RingTheory/DividedPowers/Basic.lean @@ -0,0 +1,418 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir & María-Inés de Frutos—Fernández. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir, María-Inés de Frutos—Fernández +-/ + +import Mathlib.RingTheory.PowerSeries.Basic +import Mathlib.Combinatorics.Enumerative.Bell +import Mathlib.Data.Nat.Choose.Multinomial + +/-! # Divided powers + +Let `A` be a commutative (semi)ring and `I` be an ideal of `A`. +A *divided power* structure on `I` is the datum of operations `a n ↦ dpow a n` +satisfying relations that model the intuitive formula `dpow n a = a ^ n / n !` and +collected by the structure `DividedPowers`. The list of axioms is embedded in the structure: +To avoid coercions, we rather consider `DividedPowers.dpow : ℕ → A → A`, extended by 0. + +* `DividedPowers.dpow_null` asserts that `dpow n x = 0` for `x ∉ I` + +* `DividedPowers.dpow_mem` : `dpow n x ∈ I` for `n ≠ 0` +For `x y : A` and `m n : ℕ` such that `x ∈ I` and `y ∈ I`, one has +* `DividedPowers.dpow_zero` : `dpow 0 x = 1` +* `DividedPowers.dpow_one` : `dpow 1 x = 1` +* `DividedPowers.dpow_add` : `dpow n (x + y) = +(antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y`, +this is the binomial theorem without binomial coefficients. +* `DividedPowers.dpow_mul`: `dpow n (a * x) = a ^ n * dpow n x` +* `DividedPowers.mul_dpow` : `dpow m x * dpow n x = choose (m + n) m * dpow (m + n) x` +* `DividedPowers.dpow_comp` : `dpow m (dpow n x) = uniformBell m n * dpow (m * n) x` + +* `DividedPowers.dividedPowersBot` : the trivial divided powers structure on the zero ideal + +* `DividedPowers.prod_dpow`: a product of divided powers is a multinomial coefficients +times a divided power + +* `DividedPowers.dpow_sum`: the multinomial theorem for divided powers, +without multinomial coefficients. + +* `DividedPowers.ofRingEquiv`: transfer divided powers along `RingEquiv` + +* `DividedPowers.equiv`: the equivalence `DividedPowers I ≃ DividedPowers J`, +for `e : R ≃+* S`, and `I : Ideal R`, `J : Ideal S` such that `I.map e = J` + +* `DividedPowers.exp`: the power series `Σ (dpow n a) X ^n` + +* `DividedPowers.exp_add`: its multiplicativity + +## References + +* [P. Berthelot (1974), *Cohomologie cristalline des schémas de +caractéristique $p$ > 0*][Berthelot-1974] + +* [P. Berthelot and A. Ogus (1978), *Notes on crystalline +cohomology*][BerthelotOgus-1978] + +* [N. Roby (1963), *Lois polynomes et lois formelles en théorie des +modules*][Roby-1963] + +* [N. Roby (1965), *Les algèbres à puissances dividées*][Roby-1965] + +## Discussion + +* In practice, one often has a single such structure to handle on a given ideal, +but several ideals of the same ring might be considered. +Without any explicit mention of the ideal, it is not clear whether such structures +should be provided as instances. + +* We do not provide any notation such as `a ^[n]` for `dpow a n`. + +-/ + +open Finset Nat Ideal + +section DividedPowersDefinition +/- ## Definition of divided powers -/ + +variable {A : Type*} [CommSemiring A] (I : Ideal A) + +/-- The divided power structure on an ideal I of a commutative ring A -/ +structure DividedPowers where + /-- The divided power function underlying a divided power structure -/ + dpow : ℕ → A → A + dpow_null : ∀ {n x} (_ : x ∉ I), dpow n x = 0 + dpow_zero : ∀ {x} (_ : x ∈ I), dpow 0 x = 1 + dpow_one : ∀ {x} (_ : x ∈ I), dpow 1 x = x + dpow_mem : ∀ {n x} (_ : n ≠ 0) (_ : x ∈ I), dpow n x ∈ I + dpow_add : ∀ (n) {x y} (_ : x ∈ I) (_ : y ∈ I), + dpow n (x + y) = (antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y + dpow_mul : ∀ (n) {a : A} {x} (_ : x ∈ I), + dpow n (a * x) = a ^ n * dpow n x + mul_dpow : ∀ (m n) {x} (_ : x ∈ I), + dpow m x * dpow n x = choose (m + n) m * dpow (m + n) x + dpow_comp : ∀ (m) {n x} (_ : n ≠ 0) (_ : x ∈ I), + dpow m (dpow n x) = uniformBell m n * dpow (m * n) x + +variable (A) in +/-- The canonical `DividedPowers` structure on the zero ideal -/ +def dividedPowersBot [DecidableEq A] : DividedPowers (⊥ : Ideal A) where + dpow n a := ite (a = 0 ∧ n = 0) 1 0 + dpow_null {n a} ha := by + simp only [mem_bot] at ha + dsimp + rw [if_neg] + exact not_and_of_not_left (n = 0) ha + dpow_zero {a} ha := by + rw [mem_bot.mp ha] + simp only [and_self, ite_true] + dpow_one {a} ha := by + simp [mem_bot.mp ha] + dpow_mem {n a} hn _ := by + simp only [mem_bot, ite_eq_right_iff, and_imp] + exact fun _ a ↦ False.elim (hn a) + dpow_add n a b ha hb := by + rw [mem_bot.mp ha, mem_bot.mp hb, add_zero] + simp only [true_and, ge_iff_le, tsub_eq_zero_iff_le, mul_ite, mul_one, mul_zero] + split_ifs with h + · simp [h] + · symm + apply sum_eq_zero + intro i hi + simp only [mem_antidiagonal] at hi + split_ifs with h2 h1 + · rw [h1, h2, add_zero] at hi + exfalso + exact h hi.symm + · rfl + · rfl + dpow_mul n {a x} hx := by + rw [mem_bot.mp hx] + simp only [mul_zero, true_and, mul_ite, mul_one] + by_cases hn : n = 0 + · rw [if_pos hn, hn, if_pos rfl, _root_.pow_zero] + · simp only [if_neg hn] + mul_dpow m n {x} hx := by + rw [mem_bot.mp hx] + simp only [true_and, mul_ite, mul_one, mul_zero, add_eq_zero] + by_cases hn : n = 0 + · simp only [hn, ite_true, and_true, add_zero, choose_self, cast_one] + · rw [if_neg hn, if_neg] + exact not_and_of_not_right (m = 0) hn + dpow_comp m {n a} hn ha := by + rw [mem_bot.mp ha] + simp only [true_and, ite_eq_right_iff, _root_.mul_eq_zero, mul_ite, mul_one, mul_zero] + by_cases hm: m = 0 + · simp only [hm, and_true, true_or, ite_true, uniformBell_zero_left, cast_one] + rw [if_pos] + exact fun h ↦ False.elim (hn h) + · simp only [hm, and_false, ite_false, false_or, if_neg hn] + +instance [DecidableEq A] : Inhabited (DividedPowers (⊥ : Ideal A)) := + ⟨dividedPowersBot A⟩ + +/-- The coercion from the divided powers structures to functions -/ +instance : CoeFun (DividedPowers I) fun _ ↦ ℕ → A → A := ⟨fun hI ↦ hI.dpow⟩ + +variable {I} in +@[ext] +theorem DividedPowers.ext (hI : DividedPowers I) (hI' : DividedPowers I) + (h_eq : ∀ (n : ℕ) {x : A} (_ : x ∈ I), hI.dpow n x = hI'.dpow n x) : + hI = hI' := by + obtain ⟨hI, h₀, _⟩ := hI + obtain ⟨hI', h₀', _⟩ := hI' + simp only [mk.injEq] + ext n x + by_cases hx : x ∈ I + · exact h_eq n hx + · rw [h₀ hx, h₀' hx] + +theorem DividedPowers.coe_injective : + Function.Injective (fun (h : DividedPowers I) ↦ (h : ℕ → A → A)) := fun hI hI' h ↦ by + ext n x + exact congr_fun (congr_fun h n) x + +end DividedPowersDefinition + +namespace DividedPowers + +section BasicLemmas + +/- ## Basic lemmas for divided powers -/ + +variable {A : Type*} [CommSemiring A] {I : Ideal A} {a b : A} + +/-- Variant of `DividedPowers.dpow_add` with a sum on `range (n + 1)` -/ +theorem dpow_add' (hI : DividedPowers I) (n : ℕ) (ha : a ∈ I) (hb : b ∈ I) : + hI.dpow n (a + b) = (range (n + 1)).sum fun k ↦ hI.dpow k a * hI.dpow (n - k) b := by + rw [hI.dpow_add n ha hb, sum_antidiagonal_eq_sum_range_succ_mk] + +/-- The exponential series of an element in the context of divided powers, +`Σ (dpow n a) X ^ n` -/ +def exp (hI : DividedPowers I) (a : A) : PowerSeries A := + PowerSeries.mk fun n ↦ hI.dpow n a + +/-- A more general of `DividedPowers.exp_add` -/ +theorem exp_add' (dp : ℕ → A → A) + (dp_add : ∀ n, dp n (a + b) = (antidiagonal n).sum fun k ↦ dp k.1 a * dp k.2 b) : + PowerSeries.mk (fun n ↦ dp n (a + b)) = + (PowerSeries.mk fun n ↦ dp n a) * (PowerSeries.mk fun n ↦ dp n b) := by + ext n + simp only [exp, PowerSeries.coeff_mk, PowerSeries.coeff_mul, dp_add n, + sum_antidiagonal_eq_sum_range_succ_mk] + +theorem exp_add (hI : DividedPowers I) (ha : a ∈ I) (hb : b ∈ I) : + hI.exp (a + b) = hI.exp a * hI.exp b := + exp_add' _ (fun n ↦ hI.dpow_add n ha hb) + +variable (hI : DividedPowers I) + +/- ## Rewriting lemmas -/ + +theorem dpow_smul (n : ℕ) (ha : a ∈ I) : + hI.dpow n (b • a) = b ^ n • hI.dpow n a := by + simp only [smul_eq_mul, hI.dpow_mul, ha] + +theorem dpow_mul_right (n : ℕ) (ha : a ∈ I) : + hI.dpow n (a * b) = hI.dpow n a * b ^ n := by + rw [mul_comm, hI.dpow_mul n ha, mul_comm] + +theorem dpow_smul_right (n : ℕ) (ha : a ∈ I) : + hI.dpow n (a • b) = hI.dpow n a • b ^ n := by + rw [smul_eq_mul, hI.dpow_mul_right n ha, smul_eq_mul] + +theorem factorial_mul_dpow_eq_pow (n : ℕ) (ha : a ∈ I) : + (n ! : A) * hI.dpow n a = a ^ n := by + induction n with + | zero => rw [factorial_zero, cast_one, one_mul, pow_zero, hI.dpow_zero ha] + | succ n ih => + rw [factorial_succ, mul_comm (n + 1)] + nth_rewrite 1 [← (n + 1).choose_one_right] + rw [← choose_symm_add, cast_mul, mul_assoc, + ← hI.mul_dpow n 1 ha, ← mul_assoc, ih, hI.dpow_one ha, pow_succ, mul_comm] + +theorem dpow_eval_zero {n : ℕ} (hn : n ≠ 0) : hI.dpow n 0 = 0 := by + rw [← MulZeroClass.mul_zero (0 : A), hI.dpow_mul n I.zero_mem, + zero_pow hn, zero_mul, zero_mul] + +/-- If an element of a divided power ideal is killed by multiplication +by some nonzero integer `n`, then its `n`th power is zero. + +Proposition 1.2.7 of [Berthelot-1974], part (i). -/ +theorem nilpotent_of_mem_dpIdeal {n : ℕ} (hn : n ≠ 0) (hnI : ∀ {y} (_ : y ∈ I), n • y = 0) + (hI : DividedPowers I) (ha : a ∈ I) : a ^ n = 0 := by + have h_fac : (n ! : A) * hI.dpow n a = + n • ((n - 1)! : A) * hI.dpow n a := by + rw [nsmul_eq_mul, ← cast_mul, mul_factorial_pred (Nat.pos_of_ne_zero hn)] + rw [← hI.factorial_mul_dpow_eq_pow n ha, h_fac, smul_mul_assoc] + exact hnI (I.mul_mem_left ((n - 1)! : A) (hI.dpow_mem hn ha)) + +/-- If J is another ideal of A with divided powers, +then the divided powers of I and J coincide on I • J + +[Berthelot-1974], 1.6.1 (ii) -/ +theorem coincide_on_smul {J : Ideal A} (hJ : DividedPowers J) {n : ℕ} (ha : a ∈ I • J) : + hI.dpow n a = hJ.dpow n a := by + induction ha using Submodule.smul_induction_on' generalizing n with + | smul a ha b hb => + rw [Algebra.id.smul_eq_mul, hJ.dpow_mul n hb, mul_comm a b, hI.dpow_mul n ha, ← + hJ.factorial_mul_dpow_eq_pow n hb, ← hI.factorial_mul_dpow_eq_pow n ha] + ring + | add x hx y hy hx' hy' => + rw [hI.dpow_add n (mul_le_right hx) (mul_le_right hy), + hJ.dpow_add n (mul_le_left hx) (mul_le_left hy)] + apply sum_congr rfl + intro k _ + rw [hx', hy'] + +/-- A product of divided powers is a multinomial coefficient times the divided power + +[Roby-1965], formula (III') -/ +theorem prod_dpow {ι : Type*} {s : Finset ι} (n : ι → ℕ) (ha : a ∈ I) : + (s.prod fun i ↦ hI.dpow (n i) a) = multinomial s n * hI.dpow (s.sum n) a := by + classical + induction s using Finset.induction with + | empty => + simp only [prod_empty, multinomial_empty, cast_one, sum_empty, one_mul] + rw [hI.dpow_zero ha] + | insert hi hrec => + rw [prod_insert hi, hrec, ← mul_assoc, mul_comm (hI.dpow (n _) a), + mul_assoc, mul_dpow _ _ _ ha, ← sum_insert hi, ← mul_assoc] + apply congr_arg₂ _ _ rfl + rw [multinomial_insert hi, mul_comm, cast_mul, sum_insert hi] + +-- TODO : can probably be simplified using `DividedPowers.exp` + +/-- Lemma towards `dpow_sum` when we only have partial information on a divided power ideal -/ +theorem dpow_sum' {M : Type*} [AddCommMonoid M] {I : AddSubmonoid M} (dpow : ℕ → M → A) + (dpow_zero : ∀ {x} (_ : x ∈ I), dpow 0 x = 1) + (dpow_add : ∀ (n) {x y} (_ : x ∈ I) (_ : y ∈ I), + dpow n (x + y) = (antidiagonal n).sum fun k ↦ dpow k.1 x * dpow k.2 y) + (dpow_eval_zero : ∀ {n : ℕ} (_ : n ≠ 0), dpow n 0 = 0) + {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : ι → M} (hx : ∀ i ∈ s, x i ∈ I) (n : ℕ) : + dpow n (s.sum x) = (s.sym n).sum fun k ↦ s.prod fun i ↦ dpow (Multiset.count i k) (x i) := by + simp only [sum_antidiagonal_eq_sum_range_succ_mk] at dpow_add + induction s using Finset.induction generalizing n with + | empty => + simp only [sum_empty, prod_empty, sum_const, nsmul_eq_mul, mul_one] + by_cases hn : n = 0 + · rw [hn] + rw [dpow_zero I.zero_mem] + simp only [sym_zero, card_singleton, cast_one] + · rw [dpow_eval_zero hn, eq_comm, ← cast_zero] + apply congr_arg + rw [card_eq_zero, sym_eq_empty] + exact ⟨hn, rfl⟩ + | @insert a s ha ih => + -- This should be golfable using `Finset.symInsertEquiv` + have hx' : ∀ i, i ∈ s → x i ∈ I := fun i hi ↦ hx i (mem_insert_of_mem hi) + simp_rw [sum_insert ha, + dpow_add n (hx a (mem_insert_self a s)) (I.sum_mem fun i ↦ hx' i), + sum_range, ih hx', mul_sum, sum_sigma', eq_comm] + apply sum_bij' + (fun m _ ↦ m.filterNe a) + (fun m _ ↦ m.2.fill a m.1) + (fun m hm ↦ mem_sigma.2 ⟨mem_univ _, _⟩) + (fun m hm ↦ by + simp only [succ_eq_add_one, mem_sym_iff, mem_insert, Sym.mem_fill_iff] + simp only [mem_sigma, mem_univ, mem_sym_iff, true_and] at hm + intro b + apply Or.imp (fun h ↦ h.2) (fun h ↦ hm b h)) + (fun m _ ↦ m.fill_filterNe a) + · intro m hm + simp only [mem_sigma, mem_univ, mem_sym_iff, true_and] at hm + exact Sym.filter_ne_fill a m fun a_1 ↦ ha (hm a a_1) + · intro m hm + simp only [mem_sym_iff, mem_insert] at hm + rw [prod_insert ha] + apply congr_arg₂ _ rfl + apply prod_congr rfl + intro i hi + apply congr_arg₂ _ _ rfl + conv_lhs => rw [← m.fill_filterNe a] + exact Sym.count_coe_fill_of_ne (ne_of_mem_of_not_mem hi ha) + · intro m hm + convert sym_filterNe_mem a hm + rw [erase_insert ha] + +/-- A “multinomial” theorem for divided powers — without multinomial coefficients -/ +theorem dpow_sum {ι : Type*} [DecidableEq ι] {s : Finset ι} {x : ι → A} + (hx : ∀ i ∈ s, x i ∈ I) (n : ℕ) : + hI.dpow n (s.sum x) = + (s.sym n).sum fun k ↦ s.prod fun i ↦ hI.dpow (Multiset.count i k) (x i) := + dpow_sum' hI.dpow hI.dpow_zero hI.dpow_add hI.dpow_eval_zero hx n + +end BasicLemmas + +section Equiv +/- ## Relation of divided powers with ring equivalences -/ + +variable {A B : Type*} [CommSemiring A] {I : Ideal A} [CommSemiring B] {J : Ideal B} + {e : A ≃+* B} (h : I.map e = J) + +/-- Transfer divided powers under an equivalence -/ +def ofRingEquiv (hI : DividedPowers I) : DividedPowers J where + dpow n b := e (hI.dpow n (e.symm b)) + dpow_null {n} {x} hx := by + rw [EmbeddingLike.map_eq_zero_iff, hI.dpow_null] + rwa [symm_apply_mem_of_equiv_iff, h] + dpow_zero {x} hx := by + rw [EmbeddingLike.map_eq_one_iff, hI.dpow_zero] + rwa [symm_apply_mem_of_equiv_iff, h] + dpow_one {x} hx := by + simp only + rw [dpow_one, RingEquiv.apply_symm_apply] + rwa [I.symm_apply_mem_of_equiv_iff, h] + dpow_mem {n} {x} hn hx := by + simp only + rw [← h, I.apply_mem_of_equiv_iff] + apply hI.dpow_mem hn + rwa [I.symm_apply_mem_of_equiv_iff, h] + dpow_add n {x y} hx hy := by + simp only [map_add] + rw [hI.dpow_add n (symm_apply_mem_of_equiv_iff.mpr (h ▸ hx)) + (symm_apply_mem_of_equiv_iff.mpr (h ▸ hy))] + simp only [map_sum, _root_.map_mul] + dpow_mul n {a x} hx := by + simp only [_root_.map_mul] + rw [hI.dpow_mul n (symm_apply_mem_of_equiv_iff.mpr (h ▸ hx))] + rw [_root_.map_mul, map_pow] + simp only [RingEquiv.apply_symm_apply] + mul_dpow m n {x} hx := by + simp only + rw [← _root_.map_mul, hI.mul_dpow, _root_.map_mul] + · simp only [map_natCast] + · rwa [symm_apply_mem_of_equiv_iff, h] + dpow_comp m {n x} hn hx := by + simp only [RingEquiv.symm_apply_apply] + rw [hI.dpow_comp _ hn] + · simp only [_root_.map_mul, map_natCast] + · rwa [symm_apply_mem_of_equiv_iff, h] + +@[simp] +theorem ofRingEquiv_dpow (hI : DividedPowers I) (n : ℕ) (b : B) : + (ofRingEquiv h hI).dpow n b = e (hI.dpow n (e.symm b)) := rfl + +theorem ofRingEquiv_dpow_apply (hI : DividedPowers I) (n : ℕ) (a : A) : + (ofRingEquiv h hI).dpow n (e a) = e (hI.dpow n a) := by + simp + +/-- Transfer divided powers under an equivalence (Equiv version) -/ +def equiv : DividedPowers I ≃ DividedPowers J where + toFun := ofRingEquiv h + invFun := ofRingEquiv (show map e.symm J = I by rw [← h]; exact I.map_of_equiv e) + left_inv := fun hI ↦ by ext n a; simp [ofRingEquiv] + right_inv := fun hJ ↦ by ext n b; simp [ofRingEquiv] + +theorem equiv_apply (hI : DividedPowers I) (n : ℕ) (b : B) : + (equiv h hI).dpow n b = e (hI.dpow n (e.symm b)) := rfl + +/-- Variant of `DividedPowers.equiv_apply` -/ +theorem equiv_apply' (hI : DividedPowers I) (n : ℕ) (a : A) : + (equiv h hI).dpow n (e a) = e (hI.dpow n a) := + ofRingEquiv_dpow_apply h hI n a + +end Equiv + +end DividedPowers + diff --git a/Mathlib/RingTheory/Ideal/Maps.lean b/Mathlib/RingTheory/Ideal/Maps.lean index 7a67ded065aff..e2a631a34c0a9 100644 --- a/Mathlib/RingTheory/Ideal/Maps.lean +++ b/Mathlib/RingTheory/Ideal/Maps.lean @@ -362,9 +362,20 @@ theorem comap_symm {I : Ideal R} (f : R ≃+* S) : I.comap f.symm = I.map f := /-- If `f : R ≃+* S` is a ring isomorphism and `I : Ideal R`, then `map f.symm I = comap f I`. -/ @[simp] + theorem map_symm {I : Ideal S} (f : R ≃+* S) : I.map f.symm = I.comap f := map_comap_of_equiv (RingEquiv.symm f) +@[simp] +theorem symm_apply_mem_of_equiv_iff {I : Ideal R} {f : R ≃+* S} {y : S} : + f.symm y ∈ I ↔ y ∈ I.map f := by + rw [← comap_symm, mem_comap] + +@[simp] +theorem apply_mem_of_equiv_iff {I : Ideal R} {f : R ≃+* S} {x : R} : + f x ∈ I.map f ↔ x ∈ I := by + rw [← comap_symm, Ideal.mem_comap, f.symm_apply_apply] + theorem mem_map_of_equiv {E : Type*} [EquivLike E R S] [RingEquivClass E R S] (e : E) {I : Ideal R} (y : S) : y ∈ map e I ↔ ∃ x ∈ I, e x = y := by constructor diff --git a/docs/references.bib b/docs/references.bib index 6603b6a5466fe..392120a19de1e 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -305,6 +305,31 @@ @Article{ bernstein1912 number = {1–2} } +@Book{ Berthelot-1974, + author = {Berthelot, Pierre}, + year = {1974}, + series = {Lecture {{Notes}} in {{Mathematics}}}, + number = {407}, + publisher = {sv}, + keywords = {nosource}, + file = {/Users/chambert/Zotero/storage/WHF5R9LQ/Berthelot - 1974 - + Cohomologie cristalline des schémas de caractérist.djvu}, + title = {Cohomologie Cristalline Des Schémas de Caractéristique + {$p>0$}} +} + +@Book{ BerthelotOgus-1978, + title = {Notes on Crystalline Cohomology}, + author = {Berthelot, Pierre and Ogus, Arthur}, + year = {1978}, + series = {Math. {{Notes}}}, + number = {21}, + publisher = {pup}, + keywords = {nosource}, + file = {/Users/chambert/Zotero/storage/I3FNJEDB/Berthelot et Ogus + - 1978 - Notes on crystalline cohomology.djvu} +} + @InProceedings{ beylin1996, author = "Beylin, Ilya and Dybjer, Peter", editor = "Berardi, Stefano and Coppo, Mario", @@ -3036,6 +3061,28 @@ @Misc{ RisingSea url = "https://math.stanford.edu/~vakil/216blog/" } +@Article{ Roby-1963, + title = {Lois polynomes et lois formelles en théorie des modules}, + author = {Roby, Norbert}, + year = {1963}, + journal = {Annales scientifiques de l'École Normale Supérieure}, + volume = {80}, + number = {3}, + pages = {213--348}, + doi = {10.24033/asens.1124}, + urldate = {2022-09-15}, + langid = {french} +} + +@Article{ Roby-1965, + title = {Les Algèbres à Puissances Divisées}, + author = {Roby, Norbert}, + year = {1965}, + journal = {Bulletin des Sciences Mathématiques. Deuxième Série}, + volume = {89}, + pages = {75--91} +} + @Article{ Rosenlicht_1972, author = {Rosenlicht, Maxwell}, title = {Integration in finite terms}, From b2e83338582fe358fbaad1a443a2226eeee101b1 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Thu, 21 Nov 2024 12:43:15 +0000 Subject: [PATCH 010/250] Merge master into nightly-testing --- lake-manifest.json | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 4820e7097b0d2..b696bcfdb1a73 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b3935d53ce8dad5665af7ac41f06f0f6de4d942f", + "rev": "bdd4324eda67478365b446f30c77f6bf36a40b1e", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "lean-pr-testing-6128", + "inputRev": "nightly-testing", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", From 29b03c858f01fe6ef2627bb4fce7d71e99fcc517 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Thu, 21 Nov 2024 13:54:59 +0000 Subject: [PATCH 011/250] chore: update Mathlib dependencies 2024-11-21 (#19333) 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 8fc2bedf751bd..9d7aff8b1542f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a822446d61ad7e7f5e843365c7041c326553050a", + "rev": "33d7f346440869364a2cd077bde8cebf73243aaa", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From d59b1904064942e8d375bc7091a222738e2f6539 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 21 Nov 2024 14:50:51 +0000 Subject: [PATCH 012/250] feat(NumberTheory/LSeries/Convergence): add lemma on real-valued series (#19334) This is part of the series of PRs leading to Dirichlet's Theorem on primes in AP. See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483303184) on Zulip. --- Mathlib/NumberTheory/LSeries/Convergence.lean | 18 ++++++++++++++++++ 1 file changed, 18 insertions(+) diff --git a/Mathlib/NumberTheory/LSeries/Convergence.lean b/Mathlib/NumberTheory/LSeries/Convergence.lean index c5a161fa9aae9..5310dc93a15ed 100644 --- a/Mathlib/NumberTheory/LSeries/Convergence.lean +++ b/Mathlib/NumberTheory/LSeries/Convergence.lean @@ -119,3 +119,21 @@ lemma LSeries.abscissaOfAbsConv_le_one_of_isBigO_one {f : ℕ → ℂ} (h : f =O convert abscissaOfAbsConv_le_of_isBigO_rpow (x := 0) ?_ · simp only [EReal.coe_zero, zero_add] · simpa only [Real.rpow_zero] using h + +/-- If `f` is real-valued and `x` is strictly greater than the abscissa of absolute convergence +of `f`, then the real series `∑' n, f n / n ^ x` converges. -/ +lemma LSeries.summable_real_of_abscissaOfAbsConv_lt {f : ℕ → ℝ} {x : ℝ} + (h : abscissaOfAbsConv (f ·) < x) : + Summable fun n : ℕ ↦ f n / (n : ℝ) ^ x := by + have h' : abscissaOfAbsConv (f ·) < (x : ℂ).re := by simpa only [ofReal_re] using h + have := LSeriesSummable_of_abscissaOfAbsConv_lt_re h' + rw [LSeriesSummable, show term _ _ = fun n ↦ _ from rfl] at this + conv at this => + enter [1, n] + rw [term_def, show (n : ℂ) = (n : ℝ) from rfl, ← ofReal_cpow n.cast_nonneg, ← ofReal_div, + show (0 : ℂ) = (0 : ℝ) from rfl, ← apply_ite] + rw [summable_ofReal] at this + refine this.congr_cofinite ?_ + filter_upwards [Set.Finite.compl_mem_cofinite <| Set.finite_singleton 0] with n hn + simp only [Set.mem_compl_iff, Set.mem_singleton_iff] at hn + exact if_neg hn From e2d0dfcecaa519889ccdb80b7ff61469c501d795 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Thu, 21 Nov 2024 15:34:01 +0000 Subject: [PATCH 013/250] feat: add `IsFractionRing.lift[AlgHom]_fieldRange[_eq_of_range_eq]` (#19286) Which are direct applications of `IsFractionRing.ringHom_fieldRange_eq_of_comp_eq`. Also fix some docstrings mentioning "integral domain". --- Mathlib/FieldTheory/Adjoin.lean | 15 +++++++ .../RingTheory/Localization/FractionRing.lean | 39 ++++++++++++------- 2 files changed, 41 insertions(+), 13 deletions(-) diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index e53c89cf4befb..9ccf8c539d1f9 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -1663,6 +1663,21 @@ theorem algHom_fieldRange_eq_of_comp_eq_of_range_eq refine ringHom_fieldRange_eq_of_comp_eq_of_range_eq h ?_ rw [← Algebra.adjoin_eq_ring_closure, ← hs]; rfl +variable [IsScalarTower F A K] + +/-- The image of `IsFractionRing.liftAlgHom` is the intermediate field generated by the image +of the algebra hom. -/ +theorem liftAlgHom_fieldRange (hg : Function.Injective g) : + (liftAlgHom hg : K →ₐ[F] L).fieldRange = IntermediateField.adjoin F g.range := + algHom_fieldRange_eq_of_comp_eq (by ext; simp) + +/-- The image of `IsFractionRing.liftAlgHom` is the intermediate field generated by `s`, +if the image of the algebra hom is the subalgebra generated by `s`. -/ +theorem liftAlgHom_fieldRange_eq_of_range_eq (hg : Function.Injective g) + {s : Set L} (hs : g.range = Algebra.adjoin F s) : + (liftAlgHom hg : K →ₐ[F] L).fieldRange = IntermediateField.adjoin F s := + algHom_fieldRange_eq_of_comp_eq_of_range_eq (by ext; simp) hs + end IsFractionRing set_option linter.style.longFile 1700 diff --git a/Mathlib/RingTheory/Localization/FractionRing.lean b/Mathlib/RingTheory/Localization/FractionRing.lean index fb9b40025a8df..155c5ec2966a3 100644 --- a/Mathlib/RingTheory/Localization/FractionRing.lean +++ b/Mathlib/RingTheory/Localization/FractionRing.lean @@ -39,7 +39,7 @@ variable [Algebra R S] {P : Type*} [CommRing P] variable {A : Type*} [CommRing A] (K : Type*) -- TODO: should this extend `Algebra` instead of assuming it? -/-- `IsFractionRing R K` states `K` is the field of fractions of an integral domain `R`. -/ +/-- `IsFractionRing R K` states `K` is the ring of fractions of a commutative ring `R`. -/ abbrev IsFractionRing [CommRing K] [Algebra R K] := IsLocalization (nonZeroDivisors R) K @@ -186,8 +186,8 @@ theorem mk'_eq_one_iff_eq {x : A} {y : nonZeroDivisors A} : mk' K x y = 1 ↔ x section Subfield variable (A K) in -/-- If `A` is a ring with fraction field `K`, then the subfield of `K` generated by the image of -`algebraMap A K` is equal to the whole field `K`. -/ +/-- If `A` is a commutative ring with fraction field `K`, then the subfield of `K` generated by +the image of `algebraMap A K` is equal to the whole field `K`. -/ theorem closure_range_algebraMap : Subfield.closure (Set.range (algebraMap A K)) = ⊤ := top_unique fun z _ ↦ by obtain ⟨_, _, -, rfl⟩ := div_surjective (A := A) z @@ -195,16 +195,16 @@ theorem closure_range_algebraMap : Subfield.closure (Set.range (algebraMap A K)) variable {L : Type*} [Field L] {g : A →+* L} {f : K →+* L} -/-- If `A` is a ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to `f : K →+* L`, -then the image of `f` is the field generated by the image of `g`. -/ +/-- If `A` is a commutative ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to +`f : K →+* L`, then the image of `f` is the subfield generated by the image of `g`. -/ theorem ringHom_fieldRange_eq_of_comp_eq (h : RingHom.comp f (algebraMap A K) = g) : f.fieldRange = Subfield.closure g.range := by rw [f.fieldRange_eq_map, ← closure_range_algebraMap A K, f.map_field_closure, ← Set.range_comp, ← f.coe_comp, h, g.coe_range] -/-- If `A` is a ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to `f : K →+* L`, -`s` is a set such that the image of `g` is the subring generated by `s`, -then the image of `f` is the field generated by `s`. -/ +/-- If `A` is a commutative ring with fraction field `K`, `L` is a field, `g : A →+* L` lifts to +`f : K →+* L`, `s` is a set such that the image of `g` is the subring generated by `s`, +then the image of `f` is the subfield generated by `s`. -/ theorem ringHom_fieldRange_eq_of_comp_eq_of_range_eq (h : RingHom.comp f (algebraMap A K) = g) {s : Set L} (hs : g.range = Subring.closure s) : f.fieldRange = Subfield.closure s := by rw [ringHom_fieldRange_eq_of_comp_eq h, hs] @@ -215,7 +215,7 @@ end Subfield open Function -/-- Given an integral domain `A` with field of fractions `K`, +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, we get a field hom sending `z : K` to `g x * (g y)⁻¹`, where `(x, y) : A × (NonZeroDivisors A)` are such that `z = f x * (f y)⁻¹`. -/ @@ -241,7 +241,7 @@ theorem liftAlgHom_apply : liftAlgHom hg x = lift hg x := rfl end liftAlgHom -/-- Given an integral domain `A` with field of fractions `K`, +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, the field hom induced from `K` to `L` maps `x` to `g x` for all `x : A`. -/ @@ -249,15 +249,28 @@ the field hom induced from `K` to `L` maps `x` to `g x` for all theorem lift_algebraMap (hg : Injective g) (x) : lift hg (algebraMap A K x) = g x := lift_eq _ _ -/-- Given an integral domain `A` with field of fractions `K`, +/-- The image of `IsFractionRing.lift` is the subfield generated by the image +of the ring hom. -/ +theorem lift_fieldRange (hg : Injective g) : + (lift hg : K →+* L).fieldRange = Subfield.closure g.range := + ringHom_fieldRange_eq_of_comp_eq (by ext; simp) + +/-- The image of `IsFractionRing.lift` is the subfield generated by `s`, if the image +of the ring hom is the subring generated by `s`. -/ +theorem lift_fieldRange_eq_of_range_eq (hg : Injective g) + {s : Set L} (hs : g.range = Subring.closure s) : + (lift hg : K →+* L).fieldRange = Subfield.closure s := + ringHom_fieldRange_eq_of_comp_eq_of_range_eq (by ext; simp) hs + +/-- Given a commutative ring `A` with field of fractions `K`, and an injective ring hom `g : A →+* L` where `L` is a field, field hom induced from `K` to `L` maps `f x / f y` to `g x / g y` for all `x : A, y ∈ NonZeroDivisors A`. -/ theorem lift_mk' (hg : Injective g) (x) (y : nonZeroDivisors A) : lift hg (mk' K x y) = g x / g y := by simp only [mk'_eq_div, map_div₀, lift_algebraMap] -/-- Given integral domains `A, B` with fields of fractions `K`, `L` -and an injective ring hom `j : A →+* B`, we get a field hom +/-- Given commutative rings `A, B` where `B` is an integral domain, with fraction rings `K`, `L` +and an injective ring hom `j : A →+* B`, we get a ring hom sending `z : K` to `g (j x) * (g (j y))⁻¹`, where `(x, y) : A × (NonZeroDivisors A)` are such that `z = f x * (f y)⁻¹`. -/ noncomputable def map {A B K L : Type*} [CommRing A] [CommRing B] [IsDomain B] [CommRing K] From d1c99f29771dcd3c24fa98f38a9baefa5598f548 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 21 Nov 2024 15:34:02 +0000 Subject: [PATCH 014/250] docs: `Matrix.IsTotallyUnimodular` docstring (#19338) --- .../LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 1c03274799f13..3c81f42c1c1b9 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -28,7 +28,8 @@ namespace Matrix variable {m n R : Type*} [CommRing R] -/-- Is the matrix `A` totally unimodular? -/ +/-- `A.IsTotallyUnimodular` means that every square submatrix of `A` (not necessarily contiguous) +has determinant `0` or `1` or `-1`. -/ def IsTotallyUnimodular (A : Matrix m n R) : Prop := ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, f.Injective → g.Injective → (A.submatrix f g).det = 0 ∨ From 06553ac2fd0df12b00670ebf31859d88c0bb9a52 Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Thu, 21 Nov 2024 16:01:19 +0000 Subject: [PATCH 015/250] chore: lighten tactics (#19318) Change some tactic uses to more lightweight ones (but which still automate just as much): `nlinarith` to `positivity`/`linarith`/`gcongr`, `linarith` to `norm_num`. --- Mathlib/Algebra/Order/Floor.lean | 2 +- Mathlib/Analysis/Convex/Slope.lean | 4 ++-- Mathlib/Analysis/PSeries.lean | 3 +-- Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- Mathlib/RingTheory/Ideal/Operations.lean | 4 +--- Mathlib/Topology/MetricSpace/Kuratowski.lean | 10 ++++------ 7 files changed, 11 insertions(+), 16 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor.lean b/Mathlib/Algebra/Order/Floor.lean index d3d6e7b16dc1d..0ec2dc284de96 100644 --- a/Mathlib/Algebra/Order/Floor.lean +++ b/Mathlib/Algebra/Order/Floor.lean @@ -1504,7 +1504,7 @@ theorem abs_sub_round_div_natCast_eq {m n : ℕ} : @[bound] theorem sub_half_lt_round (x : α) : x - 1 / 2 < round x := by - rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by nlinarith] + rw [round_eq x, show x - 1 / 2 = x + 1 / 2 - 1 by linarith] exact Int.sub_one_lt_floor (x + 1 / 2) @[bound] diff --git a/Mathlib/Analysis/Convex/Slope.lean b/Mathlib/Analysis/Convex/Slope.lean index 1ecc6866a18aa..51a20a60734cd 100644 --- a/Mathlib/Analysis/Convex/Slope.lean +++ b/Mathlib/Analysis/Convex/Slope.lean @@ -40,7 +40,7 @@ theorem ConvexOn.slope_mono_adjacent (hf : ConvexOn 𝕜 s f) {x y z : 𝕜} (hx field_simp [a, b, mul_comm (z - x) _] at key ⊢ rw [div_le_div_iff_of_pos_right] · linarith - · nlinarith + · positivity /-- If `f : 𝕜 → 𝕜` is concave, then for any three points `x < y < z` the slope of the secant line of `f` on `[x, y]` is greater than the slope of the secant line of `f` on `[y, z]`. -/ @@ -73,7 +73,7 @@ theorem StrictConvexOn.slope_strict_mono_adjacent (hf : StrictConvexOn 𝕜 s f) field_simp [mul_comm (z - x) _] at key ⊢ rw [div_lt_div_iff_of_pos_right] · linarith - · nlinarith + · positivity /-- If `f : 𝕜 → 𝕜` is strictly concave, then for any three points `x < y < z` the slope of the secant line of `f` on `[x, y]` is strictly greater than the slope of the secant line of `f` on diff --git a/Mathlib/Analysis/PSeries.lean b/Mathlib/Analysis/PSeries.lean index 93a65386cf5e2..38fe34d857d76 100644 --- a/Mathlib/Analysis/PSeries.lean +++ b/Mathlib/Analysis/PSeries.lean @@ -389,8 +389,7 @@ theorem sum_Ioc_inv_sq_le_sub {k n : ℕ} (hk : k ≠ 0) (h : k ≤ n) : field_simp rw [div_le_div_iff₀ _ A] · linarith - · -- Porting note: was `nlinarith` - positivity + · positivity theorem sum_Ioo_inv_sq_le (k n : ℕ) : (∑ i ∈ Ioo k n, (i ^ 2 : α)⁻¹) ≤ 2 / (k + 1) := calc diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index d05eed644c71c..71ae3fbd4d054 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -589,7 +589,7 @@ open Real Filter theorem tendsto_one_plus_div_rpow_exp (t : ℝ) : Tendsto (fun x : ℝ => (1 + t / x) ^ x) atTop (𝓝 (exp t)) := by apply ((Real.continuous_exp.tendsto _).comp (tendsto_mul_log_one_plus_div_atTop t)).congr' _ - have h₁ : (1 : ℝ) / 2 < 1 := by linarith + have h₁ : (1 : ℝ) / 2 < 1 := by norm_num have h₂ : Tendsto (fun x : ℝ => 1 + t / x) atTop (𝓝 1) := by simpa using (tendsto_inv_atTop_zero.const_mul t).const_add 1 refine (eventually_ge_of_tendsto_gt h₁ h₂).mono fun x hx => ?_ diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index b5c6b98a175d6..31cdde409beb6 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -304,7 +304,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ suffices h : p * b ^ 2 < (b ^ 2) ^ (p - 1) * b ^ 2 by apply gt_of_ge_of_gt · exact tsub_le_tsub_left (one_le_of_lt p_gt_two) ((b ^ 2) ^ (p - 1) * b ^ 2) - · have : p ≤ p * b ^ 2 := Nat.le_mul_of_pos_right _ (show 0 < b ^ 2 by nlinarith) + · have : p ≤ p * b ^ 2 := Nat.le_mul_of_pos_right _ (show 0 < b ^ 2 by positivity) exact tsub_lt_tsub_right_of_le this h suffices h : p < (b ^ 2) ^ (p - 1) by have : 4 ≤ b ^ 2 := by nlinarith diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index e18052a266b00..13dfba254995d 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -441,9 +441,7 @@ lemma sup_pow_add_le_pow_sup_pow {n m : ℕ} : (I ⊔ J) ^ (n + m) ≤ I ^ n ⊔ ((Ideal.pow_le_pow_right hn).trans le_sup_left))) · refine (Ideal.mul_le_right.trans (Ideal.mul_le_left.trans ((Ideal.pow_le_pow_right ?_).trans le_sup_right))) - simp only [Finset.mem_range, Nat.lt_succ] at hi - rw [Nat.le_sub_iff_add_le hi] - nlinarith + omega variable (I J K) diff --git a/Mathlib/Topology/MetricSpace/Kuratowski.lean b/Mathlib/Topology/MetricSpace/Kuratowski.lean index 7b7e6e47ae7f0..a0fa5566df6e3 100644 --- a/Mathlib/Topology/MetricSpace/Kuratowski.lean +++ b/Mathlib/Topology/MetricSpace/Kuratowski.lean @@ -70,12 +70,10 @@ theorem embeddingOfSubset_isometry (H : DenseRange x) : Isometry (embeddingOfSub rw [C] gcongr _ ≤ 2 * (e / 2) + dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by - have : |embeddingOfSubset x b n - embeddingOfSubset x a n| ≤ - dist (embeddingOfSubset x b) (embeddingOfSubset x a) := by - simp only [dist_eq_norm] - exact lp.norm_apply_le_norm ENNReal.top_ne_zero - (embeddingOfSubset x b - embeddingOfSubset x a) n - nlinarith + gcongr + simp only [dist_eq_norm] + exact lp.norm_apply_le_norm ENNReal.top_ne_zero + (embeddingOfSubset x b - embeddingOfSubset x a) n _ = dist (embeddingOfSubset x b) (embeddingOfSubset x a) + e := by ring simpa [dist_comm] using this From 2d53f5f766c9ed4438ad8327e84af7d714069f32 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Thu, 21 Nov 2024 17:01:14 +0000 Subject: [PATCH 016/250] =?UTF-8?q?feat(Data/Nat/Factorization/PrimePow):?= =?UTF-8?q?=20add=20equivalence=20Primes=20=C3=97=20=E2=84=95=20=E2=89=83?= =?UTF-8?q?=20PrimePowers=20(#19335)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is part of the series of PRs leading to Dirichlet's Theorem on primes in AP. See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483303184) on Zulip. --- Mathlib/Data/Nat/Factorization/PrimePow.lean | 38 ++++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/Mathlib/Data/Nat/Factorization/PrimePow.lean b/Mathlib/Data/Nat/Factorization/PrimePow.lean index 3d33d3f1dbd08..be263eea0756e 100644 --- a/Mathlib/Data/Nat/Factorization/PrimePow.lean +++ b/Mathlib/Data/Nat/Factorization/PrimePow.lean @@ -145,3 +145,41 @@ theorem Nat.mul_divisors_filter_prime_pow {a b : ℕ} (hab : a.Coprime b) : simp only [ha, hb, Finset.mem_union, Finset.mem_filter, Nat.mul_eq_zero, and_true, Ne, and_congr_left_iff, not_false_iff, Nat.mem_divisors, or_self_iff] apply hab.isPrimePow_dvd_mul + +lemma IsPrimePow.factorization_minFac_ne_zero {n : ℕ} (hn : IsPrimePow n) : + n.factorization n.minFac ≠ 0 := by + refine mt (Nat.factorization_eq_zero_iff _ _).mp ?_ + push_neg + exact ⟨n.minFac_prime hn.ne_one, n.minFac_dvd, hn.ne_zero⟩ + +/-- The canonical equivalence between pairs `(p, k)` with `p` a prime and `k : ℕ` +and the set of prime powers given by `(p, k) ↦ p^(k+1)`. -/ +def Nat.Primes.prodNatEquiv : Nat.Primes × ℕ ≃ {n : ℕ // IsPrimePow n} where + toFun pk := + ⟨pk.1 ^ (pk.2 + 1), ⟨pk.1, pk.2 + 1, prime_iff.mp pk.1.prop, pk.2.add_one_pos, rfl⟩⟩ + invFun n := + (⟨n.val.minFac, minFac_prime n.prop.ne_one⟩, n.val.factorization n.val.minFac - 1) + left_inv := fun (p, k) ↦ by + simp only [p.prop.pow_minFac k.add_one_ne_zero, Subtype.coe_eta, factorization_pow, p.prop, + Prime.factorization, Finsupp.smul_single, smul_eq_mul, mul_one, Finsupp.single_add, + Finsupp.coe_add, Pi.add_apply, Finsupp.single_eq_same, add_tsub_cancel_right] + right_inv n := by + ext1 + dsimp only + rw [sub_one_add_one n.prop.factorization_minFac_ne_zero, n.prop.minFac_pow_factorization_eq] + +@[simp] +lemma Nat.Primes.prodNatEquiv_apply (p : Nat.Primes) (k : ℕ) : + prodNatEquiv (p, k) = ⟨p ^ (k + 1), p, k + 1, prime_iff.mp p.prop, k.add_one_pos, rfl⟩ := by + rfl + +@[simp] +lemma Nat.Primes.coe_prodNatEquiv_apply (p : Nat.Primes) (k : ℕ) : + (prodNatEquiv (p, k) : ℕ) = p ^ (k + 1) := + rfl + +@[simp] +lemma Nat.Primes.prodNatEquiv_symm_apply {n : ℕ} (hn : IsPrimePow n) : + prodNatEquiv.symm ⟨n, hn⟩ = + (⟨n.minFac, minFac_prime hn.ne_one⟩, n.factorization n.minFac - 1) := + rfl From 439565f1bd4b2273a6cdcb4f4d31004ccfcb43f1 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Thu, 21 Nov 2024 21:37:19 +0000 Subject: [PATCH 017/250] refactor: use SignType to define IsTotallyUnimodular (#19345) This is a little easier than working with a three-element disjunction. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2EfromRows.20TU.20.28help.20wanted.29/near/483771954) --- .../Matrix/Determinant/TotallyUnimodular.lean | 26 +++++++++---------- 1 file changed, 12 insertions(+), 14 deletions(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index 3c81f42c1c1b9..cd7b6cae09bfd 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -5,6 +5,7 @@ Authors: Martin Dvorak, Vladimir Kolmogorov, Ivan Sergeev -/ import Mathlib.LinearAlgebra.Matrix.Determinant.Basic import Mathlib.Data.Matrix.ColumnRowPartitioned +import Mathlib.Data.Sign /-! # Totally unimodular matrices @@ -29,24 +30,20 @@ namespace Matrix variable {m n R : Type*} [CommRing R] /-- `A.IsTotallyUnimodular` means that every square submatrix of `A` (not necessarily contiguous) -has determinant `0` or `1` or `-1`. -/ +has determinant `0` or `1` or `-1`; that is, the determinant is in the range of `SignType.cast`. -/ def IsTotallyUnimodular (A : Matrix m n R) : Prop := ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, f.Injective → g.Injective → - (A.submatrix f g).det = 0 ∨ - (A.submatrix f g).det = 1 ∨ - (A.submatrix f g).det = -1 + (A.submatrix f g).det ∈ Set.range SignType.cast lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ ∀ k : ℕ, ∀ f : Fin k → m, ∀ g : Fin k → n, - (A.submatrix f g).det = 0 ∨ - (A.submatrix f g).det = 1 ∨ - (A.submatrix f g).det = -1 := by + (A.submatrix f g).det ∈ Set.range SignType.cast := by constructor <;> intro hA · intro k f g - if h : f.Injective ∧ g.Injective then - exact hA k f g h.1 h.2 - else - left + by_cases h : f.Injective ∧ g.Injective + · exact hA k f g h.1 h.2 + · refine ⟨0, ?_⟩ + rw [SignType.coe_zero, eq_comm] simp_rw [not_and_or, Function.not_injective_iff] at h obtain ⟨i, j, hfij, hij⟩ | ⟨i, j, hgij, hij⟩ := h · rw [← det_transpose, transpose_submatrix] @@ -59,10 +56,10 @@ lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) (i : m) (j : n) : - A i j = 0 ∨ A i j = 1 ∨ A i j = -1 := by + A i j ∈ Set.range SignType.cast := by let f : Fin 1 → m := (fun _ => i) let g : Fin 1 → n := (fun _ => j) - convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) <;> + convert hA 1 f g (Function.injective_of_subsingleton f) (Function.injective_of_subsingleton g) exact (det_fin_one (A.submatrix f g)).symm lemma IsTotallyUnimodular.submatrix {A : Matrix m n R} (hA : A.IsTotallyUnimodular) {k : ℕ} @@ -96,7 +93,8 @@ lemma fromRows_row0_isTotallyUnimodular_iff (A : Matrix m n R) {m' : Type*} : · exact hA k (Sum.inl ∘ f) g · if zerow : ∃ i, ∃ x', f i = Sum.inr x' then obtain ⟨i, _, _⟩ := zerow - left + use 0 + rw [eq_comm] apply det_eq_zero_of_row_eq_zero i intro simp_all From b51311315fdfbc2a8989b482a980f5277819425a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Thu, 21 Nov 2024 21:45:47 +0000 Subject: [PATCH 018/250] style(Data/Matrix/Block): define `Matrix.fromBlocks` more esthetically (#19324) --- Mathlib/Data/Matrix/Block.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Data/Matrix/Block.lean b/Mathlib/Data/Matrix/Block.lean index d572c642d2bca..cb4875d44825f 100644 --- a/Mathlib/Data/Matrix/Block.lean +++ b/Mathlib/Data/Matrix/Block.lean @@ -40,7 +40,7 @@ dimensions. -/ @[pp_nodot] def fromBlocks (A : Matrix n l α) (B : Matrix n m α) (C : Matrix o l α) (D : Matrix o m α) : Matrix (n ⊕ o) (l ⊕ m) α := - of <| Sum.elim (fun i => Sum.elim (A i) (B i)) fun i => Sum.elim (C i) (D i) + of <| Sum.elim (fun i => Sum.elim (A i) (B i)) (fun j => Sum.elim (C j) (D j)) @[simp] theorem fromBlocks_apply₁₁ (A : Matrix n l α) (B : Matrix n m α) (C : Matrix o l α) From d2f2a6b9cad40ca09cd3616fb7398a96d193b402 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Thu, 21 Nov 2024 16:10:39 -0800 Subject: [PATCH 019/250] fixes --- Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean | 2 +- Mathlib/Topology/Gluing.lean | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 5c87be9ceb0bc..df614fe6a7795 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -92,7 +92,7 @@ that the `U i`'s are open subspaces of the glued space. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] -structure GlueData extends GlueData (PresheafedSpace.{u, v, v} C) where +structure GlueData extends CategoryTheory.GlueData (PresheafedSpace.{u, v, v} C) where f_open : ∀ i j, IsOpenImmersion (f i j) attribute [instance] GlueData.f_open diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index c5385f98801ee..6c4f75721330d 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -83,7 +83,7 @@ Most of the times it would be easier to use the constructor `TopCat.GlueData.mk' conditions are stated in a less categorical way. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] -structure GlueData extends GlueData TopCat where +structure GlueData extends CategoryTheory.GlueData TopCat where f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono i j := (TopCat.mono_iff_injective _).mpr (f_open i j).isEmbedding.injective From ff67e337e215fb84ff403dd1ef4e99f51db6be8c Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 22 Nov 2024 00:49:48 +0000 Subject: [PATCH 020/250] chore(Topology/Order): move&generalize 4 lemmas (#19238) Generalize lemmas from `CompleteLinearOrder` to `ConditionallyCompleteLattice` or `ConditionallyCompleteLinearOrder`. --- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Topology/Algebra/Order/LiminfLimsup.lean | 37 ------------------- Mathlib/Topology/Order/OrderClosed.lean | 37 +++++++++++++++++++ 3 files changed, 38 insertions(+), 38 deletions(-) diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 7ab48f4222914..c9a41869344f8 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -800,7 +800,7 @@ theorem countable_meas_pos_of_disjoint_of_meas_iUnion_ne_top₀ {ι : Type*} {_ simp only [fairmeas] rfl simpa only [fairmeas_eq, posmeas_def, ← preimage_iUnion, - iUnion_Ici_eq_Ioi_of_lt_of_tendsto (0 : ℝ≥0∞) (fun n => (as_mem n).1) as_lim] + iUnion_Ici_eq_Ioi_of_lt_of_tendsto (fun n => (as_mem n).1) as_lim] rw [countable_union] refine countable_iUnion fun n => Finite.countable ?_ exact finite_const_le_meas_of_disjoint_iUnion₀ μ (as_mem n).1 As_mble As_disj Union_As_finite diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index 44fe760b02479..bc4753090a8e3 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -411,43 +411,6 @@ theorem Monotone.map_liminf_of_continuousAt {f : R → S} (f_incr : Monotone f) end Monotone -section InfiAndSupr - -open Topology - -open Filter Set - -variable [CompleteLinearOrder R] [TopologicalSpace R] [OrderTopology R] - -theorem iInf_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (x_le : ∀ i, x ≤ as i) {F : Filter ι} - [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⨅ i, as i = x := by - refine iInf_eq_of_forall_ge_of_forall_gt_exists_lt (fun i ↦ x_le i) ?_ - apply fun w x_lt_w ↦ ‹Filter.NeBot F›.nonempty_of_mem (eventually_lt_of_tendsto_lt x_lt_w as_lim) - -theorem iSup_eq_of_forall_le_of_tendsto {x : R} {as : ι → R} (le_x : ∀ i, as i ≤ x) {F : Filter ι} - [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : ⨆ i, as i = x := - iInf_eq_of_forall_le_of_tendsto (R := Rᵒᵈ) le_x as_lim - -theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto (x : R) {as : ι → R} (x_lt : ∀ i, x < as i) - {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : - ⋃ i : ι, Ici (as i) = Ioi x := by - have obs : x ∉ range as := by - intro maybe_x_is - rcases mem_range.mp maybe_x_is with ⟨i, hi⟩ - simpa only [hi, lt_self_iff_false] using x_lt i - -- Porting note: `rw at *` was too destructive. Let's only rewrite `obs` and the goal. - have := iInf_eq_of_forall_le_of_tendsto (fun i ↦ (x_lt i).le) as_lim - rw [← this] at obs - rw [← this] - exact iUnion_Ici_eq_Ioi_iInf obs - -theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto (x : R) {as : ι → R} (lt_x : ∀ i, as i < x) - {F : Filter ι} [Filter.NeBot F] (as_lim : Filter.Tendsto as F (𝓝 x)) : - ⋃ i : ι, Iic (as i) = Iio x := - iUnion_Ici_eq_Ioi_of_lt_of_tendsto (R := Rᵒᵈ) x lt_x as_lim - -end InfiAndSupr - section Indicator theorem limsup_eq_tendsto_sum_indicator_nat_atTop (s : ℕ → Set α) : diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 07a4ecd8f0010..4efa005c67f3f 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -150,6 +150,10 @@ theorem disjoint_nhds_atBot_iff : Disjoint (𝓝 a) atBot ↔ ¬IsBot a := by refine disjoint_of_disjoint_of_mem disjoint_compl_left ?_ (Iic_mem_atBot b) exact isClosed_Iic.isOpen_compl.mem_nhds hb +theorem IsLUB.range_of_tendsto {F : Filter β} [F.NeBot] (hle : ∀ i, f i ≤ a) + (hlim : Tendsto f F (𝓝 a)) : IsLUB (range f) a := + ⟨forall_mem_range.mpr hle, fun _c hc ↦ le_of_tendsto' hlim fun i ↦ hc <| mem_range_self i⟩ + end Preorder section NoBotOrder @@ -170,6 +174,23 @@ theorem not_tendsto_atBot_of_tendsto_nhds (hf : Tendsto f l (𝓝 a)) : ¬Tendst end NoBotOrder +theorem iSup_eq_of_forall_le_of_tendsto {ι : Type*} {F : Filter ι} [Filter.NeBot F] + [ConditionallyCompleteLattice α] [TopologicalSpace α] [ClosedIicTopology α] + {a : α} {f : ι → α} (hle : ∀ i, f i ≤ a) (hlim : Filter.Tendsto f F (𝓝 a)) : + ⨆ i, f i = a := + have := F.nonempty_of_neBot + (IsLUB.range_of_tendsto hle hlim).ciSup_eq + +theorem iUnion_Iic_eq_Iio_of_lt_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [ClosedIicTopology α] + {a : α} {f : ι → α} (hlt : ∀ i, f i < a) (hlim : Tendsto f F (𝓝 a)) : + ⋃ i : ι, Iic (f i) = Iio a := by + have obs : a ∉ range f := by + rw [mem_range] + rintro ⟨i, rfl⟩ + exact (hlt i).false + rw [← biUnion_range, (IsLUB.range_of_tendsto (le_of_lt <| hlt ·) hlim).biUnion_Iic_eq_Iio obs] + section LinearOrder variable [TopologicalSpace α] [LinearOrder α] [ClosedIicTopology α] [TopologicalSpace β] @@ -376,6 +397,10 @@ protected alias ⟨_, BddBelow.closure⟩ := bddBelow_closure theorem disjoint_nhds_atTop_iff : Disjoint (𝓝 a) atTop ↔ ¬IsTop a := disjoint_nhds_atBot_iff (α := αᵒᵈ) +theorem IsGLB.range_of_tendsto {F : Filter β} [F.NeBot] (hle : ∀ i, a ≤ f i) + (hlim : Tendsto f F (𝓝 a)) : IsGLB (range f) a := + IsLUB.range_of_tendsto (α := αᵒᵈ) hle hlim + end Preorder section NoTopOrder @@ -396,6 +421,18 @@ theorem not_tendsto_atTop_of_tendsto_nhds (hf : Tendsto f l (𝓝 a)) : ¬Tendst end NoTopOrder +theorem iInf_eq_of_forall_le_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLattice α] [TopologicalSpace α] [ClosedIciTopology α] + {a : α} {f : ι → α} (hle : ∀ i, a ≤ f i) (hlim : Tendsto f F (𝓝 a)) : + ⨅ i, f i = a := + iSup_eq_of_forall_le_of_tendsto (α := αᵒᵈ) hle hlim + +theorem iUnion_Ici_eq_Ioi_of_lt_of_tendsto {ι : Type*} {F : Filter ι} [F.NeBot] + [ConditionallyCompleteLinearOrder α] [TopologicalSpace α] [ClosedIciTopology α] + {a : α} {f : ι → α} (hlt : ∀ i, a < f i) (hlim : Tendsto f F (𝓝 a)) : + ⋃ i : ι, Ici (f i) = Ioi a := + iUnion_Iic_eq_Iio_of_lt_of_tendsto (α := αᵒᵈ) hlt hlim + section LinearOrder variable [TopologicalSpace α] [LinearOrder α] [ClosedIciTopology α] [TopologicalSpace β] From 1f86e43e4838ee944676f80cab98f4c66c3ed0a9 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 01:48:35 +0000 Subject: [PATCH 021/250] feat: rewrite the trailing whitespace linter in Lean (#16334) Co-authored-by: Michael Rothgang --- Mathlib/Tactic/Linter/TextBased.lean | 26 ++++++++++++++++++++------ scripts/lint-style.py | 15 +-------------- 2 files changed, 21 insertions(+), 20 deletions(-) diff --git a/Mathlib/Tactic/Linter/TextBased.lean b/Mathlib/Tactic/Linter/TextBased.lean index a1d3a9955b489..aa2c455269b43 100644 --- a/Mathlib/Tactic/Linter/TextBased.lean +++ b/Mathlib/Tactic/Linter/TextBased.lean @@ -56,6 +56,8 @@ inductive StyleError where | adaptationNote /-- A line ends with windows line endings (\r\n) instead of unix ones (\n). -/ | windowsLineEnding + /-- A line contains trailing whitespace. -/ + | trailingWhitespace deriving BEq /-- How to format style errors -/ @@ -76,6 +78,7 @@ def StyleError.errorMessage (err : StyleError) : String := match err with "Found the string \"Adaptation note:\", please use the #adaptation_note command instead" | windowsLineEnding => "This line ends with a windows line ending (\r\n): please use Unix line\ endings (\n) instead" + | trailingWhitespace => "This line ends with some whitespace: please remove this" /-- The error code for a given style error. Keep this in sync with `parse?_errorContext` below! -/ -- FUTURE: we're matching the old codes in `lint-style.py` for compatibility; @@ -83,6 +86,7 @@ def StyleError.errorMessage (err : StyleError) : String := match err with def StyleError.errorCode (err : StyleError) : String := match err with | StyleError.adaptationNote => "ERR_ADN" | StyleError.windowsLineEnding => "ERR_WIN" + | StyleError.trailingWhitespace => "ERR_TWS" /-- Context for a style error: the actual error, the line number in the file we're reading and the path to the file. -/ @@ -160,6 +164,7 @@ def parse?_errorContext (line : String) : Option ErrorContext := Id.run do -- Use default values for parameters which are ignored for comparing style exceptions. -- NB: keep this in sync with `compare` above! | "ERR_ADN" => some (StyleError.adaptationNote) + | "ERR_TWS" => some (StyleError.trailingWhitespace) | "ERR_WIN" => some (StyleError.windowsLineEnding) | _ => none match String.toNat? lineNumber with @@ -196,14 +201,24 @@ section /-- Lint on any occurrences of the string "Adaptation note:" or variants thereof. -/ def adaptationNoteLinter : TextbasedLinter := fun lines ↦ Id.run do let mut errors := Array.mkEmpty 0 - let mut lineNumber := 1 - for line in lines do + for (line, idx) in lines.zipWithIndex do -- We make this shorter to catch "Adaptation note", "adaptation note" and a missing colon. if line.containsSubstr "daptation note" then - errors := errors.push (StyleError.adaptationNote, lineNumber) - lineNumber := lineNumber + 1 + errors := errors.push (StyleError.adaptationNote, idx + 1) return (errors, none) + +/-- Lint a collection of input strings if one of them contains trailing whitespace. -/ +def trailingWhitespaceLinter : TextbasedLinter := fun lines ↦ Id.run do + let mut errors := Array.mkEmpty 0 + let mut fixedLines := lines + for (line, idx) in lines.zipWithIndex do + if line.back == ' ' then + errors := errors.push (StyleError.trailingWhitespace, idx + 1) + fixedLines := fixedLines.set! idx line.trimRight + return (errors, if errors.size > 0 then some fixedLines else none) + + /-- Whether a collection of lines consists *only* of imports, blank lines and single-line comments. In practice, this means it's an imports-only file and exempt from almost all linting. -/ def isImportsOnlyFile (lines : Array String) : Bool := @@ -215,7 +230,7 @@ end /-- All text-based linters registered in this file. -/ def allLinters : Array TextbasedLinter := #[ - adaptationNoteLinter + adaptationNoteLinter, trailingWhitespaceLinter ] @@ -257,7 +272,6 @@ def lintFile (path : FilePath) (exceptions : Array ErrorContext) : (allOutput.flatten.filter (fun e ↦ (e.find?_comparable exceptions).isNone)) return (errors, if changes_made then some changed else none) - /-- Lint a collection of modules for style violations. Print formatted errors for all unexpected style violations to standard output; correct automatically fixable style errors if configured so. diff --git a/scripts/lint-style.py b/scripts/lint-style.py index f71277d979f42..9e6353bfc09ab 100755 --- a/scripts/lint-style.py +++ b/scripts/lint-style.py @@ -39,7 +39,6 @@ ERR_IBY = 11 # isolated by ERR_IWH = 22 # isolated where ERR_SEM = 13 # the substring " ;" -ERR_TWS = 15 # trailing whitespace ERR_CLN = 16 # line starts with a colon ERR_IND = 17 # second line not correctly indented ERR_ARR = 18 # space after "←" @@ -107,15 +106,6 @@ def annotate_strings(enumerate_lines): continue yield line_nr, line, *rem, False -def line_endings_check(lines, path): - errors = [] - newlines = [] - for line_nr, line in lines: - if line.endswith(" \n"): - errors += [(ERR_TWS, line_nr, path)] - line = line.rstrip() + "\n" - newlines.append((line_nr, line)) - return errors, newlines def four_spaces_in_second_line(lines, path): # TODO: also fix the space for all lines before ":=", right now we only fix the line after @@ -248,8 +238,6 @@ def format_errors(errors): output_message(path, line_nr, "ERR_IWH", "Line is an isolated where") if errno == ERR_SEM: output_message(path, line_nr, "ERR_SEM", "Line contains a space before a semicolon") - if errno == ERR_TWS: - output_message(path, line_nr, "ERR_TWS", "Trailing whitespace detected on line") if errno == ERR_CLN: output_message(path, line_nr, "ERR_CLN", "Put : and := before line breaks, not after") if errno == ERR_IND: @@ -267,8 +255,7 @@ def lint(path, fix=False): lines = f.readlines() enum_lines = enumerate(lines, 1) newlines = enum_lines - for error_check in [line_endings_check, - four_spaces_in_second_line, + for error_check in [four_spaces_in_second_line, isolated_by_dot_semicolon_check, left_arrow_check, nonterminal_simp_check]: From 43c93172bdc359f844e718f6c24c5e958496c4ea Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Fri, 22 Nov 2024 02:00:10 +0000 Subject: [PATCH 022/250] chore(Topology/Order): rename 2 lemmas (#19261) ... to allow dot notation. --- Mathlib/Algebra/Order/Floor/Prime.lean | 2 +- Mathlib/Analysis/Asymptotics/TVS.lean | 2 +- Mathlib/Analysis/BoundedVariation.lean | 2 +- Mathlib/Analysis/Normed/Lp/lpSpace.lean | 2 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 2 +- Mathlib/Analysis/SpecificLimits/Normed.lean | 6 +++--- .../Computability/AkraBazzi/AkraBazzi.lean | 4 ++-- Mathlib/Data/Real/Pi/Irrational.lean | 4 ++-- .../Function/AEEqOfIntegral.lean | 6 +++--- Mathlib/MeasureTheory/Function/LpSpace.lean | 3 +-- .../Integral/IntegralEqImproper.lean | 4 +--- .../MeasureTheory/Measure/Portmanteau.lean | 6 +++--- .../MeasureTheory/Measure/WithDensity.lean | 2 +- Mathlib/Topology/Order/IntermediateValue.lean | 17 ++++++++-------- Mathlib/Topology/Order/OrderClosed.lean | 20 +++++++++++++++---- 15 files changed, 45 insertions(+), 37 deletions(-) diff --git a/Mathlib/Algebra/Order/Floor/Prime.lean b/Mathlib/Algebra/Order/Floor/Prime.lean index 506029de21811..8158125cc8e49 100644 --- a/Mathlib/Algebra/Order/Floor/Prime.lean +++ b/Mathlib/Algebra/Order/Floor/Prime.lean @@ -39,7 +39,7 @@ theorem exists_prime_mul_pow_div_factorial_lt_one [LinearOrderedField K] [FloorR letI := Preorder.topology K haveI : OrderTopology K := ⟨rfl⟩ ((Filter.frequently_atTop.mpr Nat.exists_infinite_primes).and_eventually - (eventually_lt_of_tendsto_lt zero_lt_one + (Filter.Tendsto.eventually_lt_const zero_lt_one (FloorSemiring.tendsto_mul_pow_div_factorial_sub_atTop a c 1))).forall_exists_of_atTop (n + 1) diff --git a/Mathlib/Analysis/Asymptotics/TVS.lean b/Mathlib/Analysis/Asymptotics/TVS.lean index f116d2b07113c..e4d49370e6a0a 100644 --- a/Mathlib/Analysis/Asymptotics/TVS.lean +++ b/Mathlib/Analysis/Asymptotics/TVS.lean @@ -125,7 +125,7 @@ lemma isLittleOTVS_one [ContinuousSMul 𝕜 E] {f : α → E} {l : Filter α} : rcases NormedField.exists_one_lt_norm 𝕜 with ⟨c, hc⟩ obtain ⟨ε, hε₀, hε⟩ : ∃ ε : ℝ≥0, 0 < ε ∧ (ε * ‖c‖₊ / r : ℝ≥0∞) < 1 := by apply Eventually.exists_gt - refine eventually_lt_of_tendsto_lt zero_lt_one <| Continuous.tendsto' ?_ _ _ (by simp) + refine Continuous.tendsto' ?_ _ _ (by simp) |>.eventually_lt_const zero_lt_one fun_prop (disch := intros; first | apply ENNReal.coe_ne_top | positivity) filter_upwards [hr ε hε₀.ne'] with x hx refine mem_of_egauge_lt_one hUb (hx.trans_lt ?_) diff --git a/Mathlib/Analysis/BoundedVariation.lean b/Mathlib/Analysis/BoundedVariation.lean index 4f281d9bfca6d..4a2fafb5a8000 100644 --- a/Mathlib/Analysis/BoundedVariation.lean +++ b/Mathlib/Analysis/BoundedVariation.lean @@ -180,7 +180,7 @@ theorem lowerSemicontinuous_aux {ι : Type*} {F : ι → α → E} {p : Filter (𝓝 (∑ i ∈ Finset.range n, edist (f (u (i + 1))) (f (u i)))) := by apply tendsto_finset_sum exact fun i _ => Tendsto.edist (Ffs (u i.succ) (us i.succ)) (Ffs (u i) (us i)) - exact (eventually_gt_of_tendsto_gt hlt this).mono fun i h => h.trans_le (sum_le (F i) n um us) + exact (this.eventually_const_lt hlt).mono fun i h => h.trans_le (sum_le (F i) n um us) /-- The map `(eVariationOn · s)` is lower semicontinuous for pointwise convergence *on `s`*. Pointwise convergence on `s` is encoded here as uniform convergence on the family consisting of the diff --git a/Mathlib/Analysis/Normed/Lp/lpSpace.lean b/Mathlib/Analysis/Normed/Lp/lpSpace.lean index 4797943bc7ec5..b2a36ecac89b5 100644 --- a/Mathlib/Analysis/Normed/Lp/lpSpace.lean +++ b/Mathlib/Analysis/Normed/Lp/lpSpace.lean @@ -183,7 +183,7 @@ theorem of_exponent_ge {p q : ℝ≥0∞} {f : ∀ i, E i} (hfq : Memℓp f q) ( have hf' := hfq.summable hq refine .of_norm_bounded_eventually _ hf' (@Set.Finite.subset _ { i | 1 ≤ ‖f i‖ } ?_ _ ?_) · have H : { x : α | 1 ≤ ‖f x‖ ^ q.toReal }.Finite := by - simpa using eventually_lt_of_tendsto_lt (by norm_num) hf'.tendsto_cofinite_zero + simpa using hf'.tendsto_cofinite_zero.eventually_lt_const (by norm_num) exact H.subset fun i hi => Real.one_le_rpow hi hq.le · show ∀ i, ¬|‖f i‖ ^ p.toReal| ≤ ‖f i‖ ^ q.toReal → 1 ≤ ‖f i‖ intro i hi diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 71ae3fbd4d054..49c7cb8a86df9 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -592,7 +592,7 @@ theorem tendsto_one_plus_div_rpow_exp (t : ℝ) : have h₁ : (1 : ℝ) / 2 < 1 := by norm_num have h₂ : Tendsto (fun x : ℝ => 1 + t / x) atTop (𝓝 1) := by simpa using (tendsto_inv_atTop_zero.const_mul t).const_add 1 - refine (eventually_ge_of_tendsto_gt h₁ h₂).mono fun x hx => ?_ + refine (h₂.eventually_const_le h₁).mono fun x hx => ?_ have hx' : 0 < 1 + t / x := by linarith simp [mul_comm x, exp_mul, exp_log hx'] diff --git a/Mathlib/Analysis/SpecificLimits/Normed.lean b/Mathlib/Analysis/SpecificLimits/Normed.lean index 69992fdb4a113..cf5f9286f5bd2 100644 --- a/Mathlib/Analysis/SpecificLimits/Normed.lean +++ b/Mathlib/Analysis/SpecificLimits/Normed.lean @@ -601,7 +601,7 @@ theorem summable_of_ratio_test_tendsto_lt_one {α : Type*} [NormedAddCommGroup (h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) : Summable f := by rcases exists_between hl₁ with ⟨r, hr₀, hr₁⟩ refine summable_of_ratio_norm_eventually_le hr₁ ?_ - filter_upwards [eventually_le_of_tendsto_lt hr₀ h, hf] with _ _ h₁ + filter_upwards [h.eventually_le_const hr₀, hf] with _ _ h₁ rwa [← div_le_iff₀ (norm_pos_iff.mpr h₁)] theorem not_summable_of_ratio_norm_eventually_ge {α : Type*} [SeminormedAddCommGroup α] {f : ℕ → α} @@ -629,12 +629,12 @@ theorem not_summable_of_ratio_test_tendsto_gt_one {α : Type*} [SeminormedAddCom {f : ℕ → α} {l : ℝ} (hl : 1 < l) (h : Tendsto (fun n ↦ ‖f (n + 1)‖ / ‖f n‖) atTop (𝓝 l)) : ¬Summable f := by have key : ∀ᶠ n in atTop, ‖f n‖ ≠ 0 := by - filter_upwards [eventually_ge_of_tendsto_gt hl h] with _ hn hc + filter_upwards [h.eventually_const_le hl] with _ hn hc rw [hc, _root_.div_zero] at hn linarith rcases exists_between hl with ⟨r, hr₀, hr₁⟩ refine not_summable_of_ratio_norm_eventually_ge hr₀ key.frequently ?_ - filter_upwards [eventually_ge_of_tendsto_gt hr₁ h, key] with _ _ h₁ + filter_upwards [h.eventually_const_le hr₁, key] with _ _ h₁ rwa [← le_div_iff₀ (lt_of_le_of_ne (norm_nonneg _) h₁.symm)] section NormedDivisionRing diff --git a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean index 9cfbcd713b62a..7cd217fadb8cf 100644 --- a/Mathlib/Computability/AkraBazzi/AkraBazzi.lean +++ b/Mathlib/Computability/AkraBazzi/AkraBazzi.lean @@ -544,9 +544,9 @@ lemma tendsto_atTop_sumCoeffsExp : Tendsto (fun (p : ℝ) => ∑ i, a i * (b i) lemma one_mem_range_sumCoeffsExp : 1 ∈ Set.range (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := by refine mem_range_of_exists_le_of_exists_ge R.continuous_sumCoeffsExp ?le_one ?ge_one case le_one => - exact Eventually.exists <| eventually_le_of_tendsto_lt zero_lt_one R.tendsto_zero_sumCoeffsExp + exact R.tendsto_zero_sumCoeffsExp.eventually_le_const zero_lt_one |>.exists case ge_one => - exact Eventually.exists <| R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _ + exact R.tendsto_atTop_sumCoeffsExp.eventually_ge_atTop _ |>.exists /-- The function x ↦ ∑ a_i b_i^x is injective. This implies the uniqueness of `p`. -/ lemma injective_sumCoeffsExp : Function.Injective (fun (p : ℝ) => ∑ i, a i * (b i) ^ p) := diff --git a/Mathlib/Data/Real/Pi/Irrational.lean b/Mathlib/Data/Real/Pi/Irrational.lean index e3884b774698d..afebb88a105b2 100644 --- a/Mathlib/Data/Real/Pi/Irrational.lean +++ b/Mathlib/Data/Real/Pi/Irrational.lean @@ -283,8 +283,8 @@ private lemma not_irrational_exists_rep {x : ℝ} : rwa [lt_div_iff₀ (by positivity), zero_mul] at this have k (n : ℕ) : 0 < (a : ℝ) ^ (2 * n + 1) / n ! := by positivity have j : ∀ᶠ n : ℕ in atTop, (a : ℝ) ^ (2 * n + 1) / n ! * I n (π / 2) < 1 := by - have := eventually_lt_of_tendsto_lt (show (0 : ℝ) < 1 / 2 by norm_num) - (tendsto_pow_div_factorial_at_top_aux a) + have := (tendsto_pow_div_factorial_at_top_aux a).eventually_lt_const + (show (0 : ℝ) < 1 / 2 by norm_num) filter_upwards [this] with n hn rw [lt_div_iff₀ (zero_lt_two : (0 : ℝ) < 2)] at hn exact hn.trans_le' (mul_le_mul_of_nonneg_left (I_le _) (by positivity)) diff --git a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean index ce393e78fea5c..10f195e4aacff 100644 --- a/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean +++ b/Mathlib/MeasureTheory/Function/AEEqOfIntegral.lean @@ -193,12 +193,12 @@ theorem ae_le_of_forall_setLIntegral_le_of_sigmaFinite₀ [SigmaFinite μ] have : Tendsto (fun n => g x + u n) atTop (𝓝 (g x + (0 : ℝ≥0))) := tendsto_const_nhds.add (ENNReal.tendsto_coe.2 u_lim) simp only [ENNReal.coe_zero, add_zero] at this - exact eventually_le_of_tendsto_lt hx this + exact this.eventually_le_const hx have L2 : ∀ᶠ n : ℕ in (atTop : Filter ℕ), g x ≤ (n : ℝ≥0) := - haveI : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by + have : Tendsto (fun n : ℕ => ((n : ℝ≥0) : ℝ≥0∞)) atTop (𝓝 ∞) := by simp only [ENNReal.coe_natCast] exact ENNReal.tendsto_nat_nhds_top - eventually_ge_of_tendsto_gt (hx.trans_le le_top) this + this.eventually_const_le (hx.trans_le le_top) apply Set.mem_iUnion.2 exact ((L1.and L2).and (eventually_mem_spanningSets μ x)).exists refine le_antisymm ?_ bot_le diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 01ccd42815470..5b9c9290dac61 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -673,8 +673,7 @@ theorem exists_eLpNorm_indicator_le (hp : p ≠ ∞) (c : E) {ε : ℝ≥0∞} ( convert (NNReal.continuousAt_rpow_const (Or.inr hp₀')).tendsto.const_mul _ simp [hp₀''.ne'] have hε' : 0 < ε := hε.bot_lt - obtain ⟨δ, hδ, hδε'⟩ := - NNReal.nhds_zero_basis.eventually_iff.mp (eventually_le_of_tendsto_lt hε' this) + obtain ⟨δ, hδ, hδε'⟩ := NNReal.nhds_zero_basis.eventually_iff.mp (this.eventually_le_const hε') obtain ⟨η, hη, hηδ⟩ := exists_between hδ refine ⟨η, hη, ?_⟩ rw [← ENNReal.coe_rpow_of_nonneg _ hp₀', ← ENNReal.coe_mul] diff --git a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean index e2ec1a0114d79..9678b171ff3fc 100644 --- a/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean +++ b/Mathlib/MeasureTheory/Integral/IntegralEqImproper.lean @@ -388,9 +388,7 @@ theorem AECover.iSup_lintegral_eq_of_countably_generated [Nonempty ι] [l.NeBot] have := hφ.lintegral_tendsto_of_countably_generated hfm refine ciSup_eq_of_forall_le_of_forall_lt_exists_gt (fun i => lintegral_mono' Measure.restrict_le_self le_rfl) fun w hw => ?_ - rcases exists_between hw with ⟨m, hm₁, hm₂⟩ - rcases (eventually_ge_of_tendsto_gt hm₂ this).exists with ⟨i, hi⟩ - exact ⟨i, lt_of_lt_of_le hm₁ hi⟩ + exact (this.eventually_const_lt hw).exists end Lintegral diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 4d82e897d1771..72fe3305828be 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -278,14 +278,14 @@ theorem FiniteMeasure.limsup_measure_closed_le_of_tendsto {Ω ι : Type*} {L : F HasOuterApproxClosed.tendsto_lintegral_apprSeq F_closed (μ : Measure Ω) have room₁ : (μ : Measure Ω) F < (μ : Measure Ω) F + ε / 2 := ENNReal.lt_add_right (measure_lt_top (μ : Measure Ω) F).ne ε_pos' - obtain ⟨M, hM⟩ := eventually_atTop.mp <| eventually_lt_of_tendsto_lt room₁ key₁ + obtain ⟨M, hM⟩ := eventually_atTop.mp <| key₁.eventually_lt_const room₁ have key₂ := FiniteMeasure.tendsto_iff_forall_lintegral_tendsto.mp μs_lim (fs M) have room₂ : (lintegral (μ : Measure Ω) fun a ↦ fs M a) < (lintegral (μ : Measure Ω) fun a ↦ fs M a) + ε / 2 := ENNReal.lt_add_right (ne_of_lt ((fs M).lintegral_lt_top_of_nnreal _)) ε_pos' - have ev_near := Eventually.mono (eventually_lt_of_tendsto_lt room₂ key₂) fun n ↦ le_of_lt - have ev_near' := Eventually.mono ev_near + have ev_near := key₂.eventually_le_const room₂ + have ev_near' := ev_near.mono (fun n ↦ le_trans (HasOuterApproxClosed.measure_le_lintegral F_closed (μs n) M)) apply (Filter.limsup_le_limsup ev_near').trans rw [limsup_const] diff --git a/Mathlib/MeasureTheory/Measure/WithDensity.lean b/Mathlib/MeasureTheory/Measure/WithDensity.lean index d343307b3e3fc..f8c9285e8bbbf 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensity.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensity.lean @@ -669,7 +669,7 @@ lemma IsLocallyFiniteMeasure.withDensity_coe {f : α → ℝ≥0} (hf : Continuo IsLocallyFiniteMeasure (μ.withDensity fun x ↦ f x) := by refine ⟨fun x ↦ ?_⟩ rcases (μ.finiteAt_nhds x).exists_mem_basis ((nhds_basis_opens' x).restrict_subset - (eventually_le_of_tendsto_lt (lt_add_one _) (hf.tendsto x))) with ⟨U, ⟨⟨hUx, hUo⟩, hUf⟩, hμU⟩ + ((hf.tendsto x).eventually_le_const (lt_add_one _))) with ⟨U, ⟨⟨hUx, hUo⟩, hUf⟩, hμU⟩ refine ⟨U, hUx, ?_⟩ rw [withDensity_apply _ hUo.measurableSet] exact setLIntegral_lt_top_of_bddAbove hμU.ne ⟨f x + 1, forall_mem_image.2 hUf⟩ diff --git a/Mathlib/Topology/Order/IntermediateValue.lean b/Mathlib/Topology/Order/IntermediateValue.lean index 81446ea8fdf08..ef73d422348e1 100644 --- a/Mathlib/Topology/Order/IntermediateValue.lean +++ b/Mathlib/Topology/Order/IntermediateValue.lean @@ -122,21 +122,20 @@ theorem IsPreconnected.intermediate_value {s : Set X} (hs : IsPreconnected s) {a theorem IsPreconnected.intermediate_value_Ico {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht : Tendsto f l (𝓝 v)) : Ico (f a) v ⊆ f '' s := fun _ h => - hs.intermediate_value₂_eventually₁ ha hl hf continuousOn_const h.1 - (eventually_ge_of_tendsto_gt h.2 ht) + hs.intermediate_value₂_eventually₁ ha hl hf continuousOn_const h.1 (ht.eventually_const_le h.2) theorem IsPreconnected.intermediate_value_Ioc {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht : Tendsto f l (𝓝 v)) : Ioc v (f a) ⊆ f '' s := fun _ h => (hs.intermediate_value₂_eventually₁ ha hl continuousOn_const hf h.2 - (eventually_le_of_tendsto_lt h.1 ht)).imp fun _ h => h.imp_right Eq.symm + (ht.eventually_le_const h.1)).imp fun _ h => h.imp_right Eq.symm theorem IsPreconnected.intermediate_value_Ioo {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v₁ v₂ : α} (ht₁ : Tendsto f l₁ (𝓝 v₁)) (ht₂ : Tendsto f l₂ (𝓝 v₂)) : Ioo v₁ v₂ ⊆ f '' s := fun _ h => hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const - (eventually_le_of_tendsto_lt h.1 ht₁) (eventually_ge_of_tendsto_gt h.2 ht₂) + (ht₁.eventually_le_const h.1) (ht₂.eventually_const_le h.2) theorem IsPreconnected.intermediate_value_Ici {s : Set X} (hs : IsPreconnected s) {a : X} {l : Filter X} (ha : a ∈ s) [NeBot l] (hl : l ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) @@ -153,19 +152,19 @@ theorem IsPreconnected.intermediate_value_Ioi {s : Set X} (hs : IsPreconnected s [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht₁ : Tendsto f l₁ (𝓝 v)) (ht₂ : Tendsto f l₂ atTop) : Ioi v ⊆ f '' s := fun y h => hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const - (eventually_le_of_tendsto_lt h ht₁) (tendsto_atTop.1 ht₂ y) + (ht₁.eventually_le_const h) (ht₂.eventually_ge_atTop y) theorem IsPreconnected.intermediate_value_Iio {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) {v : α} (ht₁ : Tendsto f l₁ atBot) (ht₂ : Tendsto f l₂ (𝓝 v)) : Iio v ⊆ f '' s := fun y h => - hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (tendsto_atBot.1 ht₁ y) - (eventually_ge_of_tendsto_gt h ht₂) + hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (ht₁.eventually_le_atBot y) + (ht₂.eventually_const_le h) theorem IsPreconnected.intermediate_value_Iii {s : Set X} (hs : IsPreconnected s) {l₁ l₂ : Filter X} [NeBot l₁] [NeBot l₂] (hl₁ : l₁ ≤ 𝓟 s) (hl₂ : l₂ ≤ 𝓟 s) {f : X → α} (hf : ContinuousOn f s) (ht₁ : Tendsto f l₁ atBot) (ht₂ : Tendsto f l₂ atTop) : univ ⊆ f '' s := fun y _ => - hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (tendsto_atBot.1 ht₁ y) - (tendsto_atTop.1 ht₂ y) + hs.intermediate_value₂_eventually₂ hl₁ hl₂ hf continuousOn_const (ht₁.eventually_le_atBot y) + (ht₂.eventually_ge_atTop y) /-- **Intermediate Value Theorem** for continuous functions on connected spaces. -/ theorem intermediate_value_univ [PreconnectedSpace X] (a b : X) {f : X → α} (hf : Continuous f) : diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 4efa005c67f3f..7ded140f87a07 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -213,14 +213,20 @@ theorem Ici_mem_nhds (h : a < b) : Ici a ∈ 𝓝 b := theorem eventually_ge_nhds (hab : b < a) : ∀ᶠ x in 𝓝 a, b ≤ x := Ici_mem_nhds hab -theorem eventually_gt_of_tendsto_gt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) +theorem Filter.Tendsto.eventually_const_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) (h : Filter.Tendsto f l (𝓝 v)) : ∀ᶠ a in l, u < f a := h.eventually <| eventually_gt_nhds hv -theorem eventually_ge_of_tendsto_gt {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) +@[deprecated (since := "2024-11-17")] +alias eventually_gt_of_tendsto_gt := Filter.Tendsto.eventually_const_lt + +theorem Filter.Tendsto.eventually_const_le {l : Filter γ} {f : γ → α} {u v : α} (hv : u < v) (h : Tendsto f l (𝓝 v)) : ∀ᶠ a in l, u ≤ f a := h.eventually <| eventually_ge_nhds hv +@[deprecated (since := "2024-11-17")] +alias eventually_ge_of_tendsto_gt := Filter.Tendsto.eventually_const_le + protected theorem Dense.exists_gt [NoMaxOrder α] {s : Set α} (hs : Dense s) (x : α) : ∃ y ∈ s, x < y := hs.exists_mem_open isOpen_Ioi (exists_gt x) @@ -451,14 +457,20 @@ theorem Iic_mem_nhds (h : a < b) : Iic b ∈ 𝓝 a := theorem eventually_le_nhds (hab : a < b) : ∀ᶠ x in 𝓝 a, x ≤ b := Iic_mem_nhds hab -theorem eventually_lt_of_tendsto_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) +theorem Filter.Tendsto.eventually_lt_const {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) (h : Filter.Tendsto f l (𝓝 v)) : ∀ᶠ a in l, f a < u := h.eventually <| eventually_lt_nhds hv -theorem eventually_le_of_tendsto_lt {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) +@[deprecated (since := "2024-11-17")] +alias eventually_lt_of_tendsto_lt := Filter.Tendsto.eventually_lt_const + +theorem Filter.Tendsto.eventually_le_const {l : Filter γ} {f : γ → α} {u v : α} (hv : v < u) (h : Tendsto f l (𝓝 v)) : ∀ᶠ a in l, f a ≤ u := h.eventually <| eventually_le_nhds hv +@[deprecated (since := "2024-11-17")] +alias eventually_le_of_tendsto_lt := Filter.Tendsto.eventually_le_const + protected theorem Dense.exists_lt [NoMinOrder α] {s : Set α} (hs : Dense s) (x : α) : ∃ y ∈ s, y < x := hs.orderDual.exists_gt x From fe0e8bcc2ddc54118bbd2abdc474dd1d2e3076f3 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Fri, 22 Nov 2024 04:40:29 +0000 Subject: [PATCH 023/250] feat(NumberTheory/LSeries/PrimesInAP): new file (#19344) This PR creates a new file `Mathlib.NumberTheory.LSeries.PrimesInAP` that will eventually contain a proof of **Dirichlet's Theorem on primes in arithmetic progression**. As a first step, we provide two lemmas on re-writing sums (or products) over a function supported in prime powers. (We also add an API lemma for `Multipliable`/`Summable` that seems to be missing.) See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483714489) on Zulip. --- Mathlib.lean | 1 + Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 67 +++++++++++++++++++ .../Algebra/InfiniteSum/Constructions.lean | 5 ++ 3 files changed, 73 insertions(+) create mode 100644 Mathlib/NumberTheory/LSeries/PrimesInAP.lean diff --git a/Mathlib.lean b/Mathlib.lean index e6b62736993a2..5e0aee1763797 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3770,6 +3770,7 @@ import Mathlib.NumberTheory.LSeries.Linearity import Mathlib.NumberTheory.LSeries.MellinEqDirichlet import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.NumberTheory.LSeries.Positivity +import Mathlib.NumberTheory.LSeries.PrimesInAP import Mathlib.NumberTheory.LSeries.RiemannZeta import Mathlib.NumberTheory.LSeries.ZMod import Mathlib.NumberTheory.LegendreSymbol.AddCharacter diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean new file mode 100644 index 0000000000000..0c1fc1a1ebf26 --- /dev/null +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -0,0 +1,67 @@ +/- +Copyright (c) 2024 Michael Stoll. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Michael Stoll +-/ +import Mathlib.Data.Nat.Factorization.PrimePow +import Mathlib.Topology.Algebra.InfiniteSum.Constructions +import Mathlib.Topology.Algebra.InfiniteSum.NatInt + +/-! +# Dirichlet's Theorem on primes in arithmetic progression + +The goal of this file is to prove **Dirichlet's Theorem**: If `q` is a positive natural number +and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` such that +`(p : ZMod q) = a`. + +This will be done in the following steps: + +1. Some auxiliary lemmas on infinite products and sums +2. (TODO) Results on the von Mangoldt function restricted to a residue class +3. (TODO) Results on its L-series +4. (TODO) Derivation of Dirichlet's Theorem +-/ + +/-! +### Auxiliary statements + +An infinite product or sum over a function supported in prime powers can be written +as an iterated product or sum over primes and natural numbers. +-/ + +section auxiliary + +variable {α β γ : Type*} [CommGroup α] [UniformSpace α] [UniformGroup α] [CompleteSpace α] + [T0Space α] + +open Nat.Primes in +@[to_additive tsum_eq_tsum_primes_of_support_subset_prime_powers] +lemma tprod_eq_tprod_primes_of_mulSupport_subset_prime_powers {f : ℕ → α} + (hfm : Multipliable f) (hf : Function.mulSupport f ⊆ {n | IsPrimePow n}) : + ∏' n : ℕ, f n = ∏' (p : Nat.Primes) (k : ℕ), f (p ^ (k + 1)) := by + have hfm' : Multipliable fun pk : Nat.Primes × ℕ ↦ f (pk.fst ^ (pk.snd + 1)) := + prodNatEquiv.symm.multipliable_iff.mp <| by + simpa only [← coe_prodNatEquiv_apply, Prod.eta, Function.comp_def, Equiv.apply_symm_apply] + using hfm.subtype _ + simp only [← tprod_subtype_eq_of_mulSupport_subset hf, Set.coe_setOf, ← prodNatEquiv.tprod_eq, + ← tprod_prod hfm'] + refine tprod_congr fun (p, k) ↦ congrArg f <| coe_prodNatEquiv_apply .. + +@[to_additive tsum_eq_tsum_primes_add_tsum_primes_of_support_subset_prime_powers] +lemma tprod_eq_tprod_primes_mul_tprod_primes_of_mulSupport_subset_prime_powers {f : ℕ → α} + (hfm : Multipliable f) (hf : Function.mulSupport f ⊆ {n | IsPrimePow n}) : + ∏' n : ℕ, f n = (∏' p : Nat.Primes, f p) * ∏' (p : Nat.Primes) (k : ℕ), f (p ^ (k + 2)) := by + rw [tprod_eq_tprod_primes_of_mulSupport_subset_prime_powers hfm hf] + have hfs' (p : Nat.Primes) : Multipliable fun k : ℕ ↦ f (p ^ (k + 1)) := + hfm.comp_injective <| (strictMono_nat_of_lt_succ + fun k ↦ pow_lt_pow_right₀ p.prop.one_lt <| lt_add_one (k + 1)).injective + conv_lhs => + enter [1, p]; rw [tprod_eq_zero_mul (hfs' p), zero_add, pow_one] + enter [2, 1, k]; rw [add_assoc, one_add_one_eq_two] + exact tprod_mul (Multipliable.subtype hfm _) <| + Multipliable.prod (f := fun (pk : Nat.Primes × ℕ) ↦ f (pk.1 ^ (pk.2 + 2))) <| + hfm.comp_injective <| Subtype.val_injective |>.comp + Nat.Primes.prodNatEquiv.injective |>.comp <| + Function.Injective.prodMap (fun ⦃_ _⦄ a ↦ a) <| add_left_injective 1 + +end auxiliary diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index 47d09ed95c689..8b1f1a092440e 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -192,6 +192,11 @@ theorem Multipliable.prod_factor {f : β × γ → α} (h : Multipliable f) (b : Multipliable fun c ↦ f (b, c) := h.comp_injective fun _ _ h ↦ (Prod.ext_iff.1 h).2 +@[to_additive Summable.prod] +lemma Multipliable.prod {f : β × γ → α} (h : Multipliable f) : + Multipliable fun b ↦ ∏' c, f (b, c) := + ((Equiv.sigmaEquivProd β γ).multipliable_iff.mpr h).sigma + @[to_additive] lemma HasProd.tprod_fiberwise [T2Space α] {f : β → α} {a : α} (hf : HasProd f a) (g : β → γ) : HasProd (fun c : γ ↦ ∏' b : g ⁻¹' {c}, f b) a := From 981dca5c7e4eea030c1ea4e06efb682c3e784d9c Mon Sep 17 00:00:00 2001 From: sgouezel Date: Fri, 22 Nov 2024 05:01:05 +0000 Subject: [PATCH 024/250] chore(SchwartzSpace): clean up white space (#19336) Definitions in term mode using `mkCLM` are indented by 6 spaces, but formulating things with `refine` drops it to a more reasonable 2 spaces. Absolutely no math change, just unindenting. --- .../Analysis/Distribution/SchwartzSpace.lean | 234 +++++++++--------- 1 file changed, 114 insertions(+), 120 deletions(-) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 6debdf4d51000..0c1edc7031b70 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -764,53 +764,51 @@ variable [NormedAddCommGroup G] [NormedSpace ℝ G] /-- The map `f ↦ (x ↦ B (f x) (g x))` as a continuous `𝕜`-linear map on Schwartz space, where `B` is a continuous `𝕜`-linear map and `g` is a function of temperate growth. -/ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemperateGrowth) : - 𝓢(D, E) →L[ℝ] 𝓢(D, G) := + 𝓢(D, E) →L[ℝ] 𝓢(D, G) := by -- Todo (after port): generalize to `B : E →L[𝕜] F →L[𝕜] G` and `𝕜`-linear - mkCLM - (fun f x => B (f x) (g x)) + refine mkCLM (fun f x => B (f x) (g x)) (fun _ _ _ => by simp only [map_add, add_left_inj, Pi.add_apply, eq_self_iff_true, ContinuousLinearMap.add_apply]) (fun _ _ _ => by simp only [smul_apply, map_smul, ContinuousLinearMap.coe_smul', Pi.smul_apply, RingHom.id_apply]) - (fun f => (B.isBoundedBilinearMap.contDiff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) - (by - rintro ⟨k, n⟩ - rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ - use - Finset.Iic (l + k, n), ‖B‖ * ((n : ℝ) + (1 : ℝ)) * n.choose (n / 2) * (C * 2 ^ (l + k)), - by positivity - intro f x - have hxk : 0 ≤ ‖x‖ ^ k := by positivity - have hnorm_mul := - ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top - refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ - move_mul [← ‖B‖] - simp_rw [mul_assoc ‖B‖] - gcongr _ * ?_ - rw [Finset.mul_sum] - have : (∑ _x ∈ Finset.range (n + 1), (1 : ℝ)) = n + 1 := by simp - simp_rw [mul_assoc ((n : ℝ) + 1)] - rw [← this, Finset.sum_mul] - refine Finset.sum_le_sum fun i hi => ?_ - simp only [one_mul] - move_mul [(Nat.choose n i : ℝ), (Nat.choose n (n / 2) : ℝ)] - gcongr ?_ * ?_ - swap - · norm_cast - exact i.choose_le_middle n - specialize hgrowth (n - i) (by simp only [tsub_le_self]) x - refine le_trans (mul_le_mul_of_nonneg_left hgrowth (by positivity)) ?_ - move_mul [C] - gcongr ?_ * C - rw [Finset.mem_range_succ_iff] at hi - change i ≤ (l + k, n).snd at hi - refine le_trans ?_ (one_add_le_sup_seminorm_apply le_rfl hi f x) - rw [pow_add] - move_mul [(1 + ‖x‖) ^ l] - gcongr - simp) + (fun f => (B.isBoundedBilinearMap.contDiff.restrict_scalars ℝ).comp (f.smooth'.prod hg.1)) ?_ + rintro ⟨k, n⟩ + rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ + use + Finset.Iic (l + k, n), ‖B‖ * ((n : ℝ) + (1 : ℝ)) * n.choose (n / 2) * (C * 2 ^ (l + k)), + by positivity + intro f x + have hxk : 0 ≤ ‖x‖ ^ k := by positivity + have hnorm_mul := + ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top + refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ + move_mul [← ‖B‖] + simp_rw [mul_assoc ‖B‖] + gcongr _ * ?_ + rw [Finset.mul_sum] + have : (∑ _x ∈ Finset.range (n + 1), (1 : ℝ)) = n + 1 := by simp + simp_rw [mul_assoc ((n : ℝ) + 1)] + rw [← this, Finset.sum_mul] + refine Finset.sum_le_sum fun i hi => ?_ + simp only [one_mul] + move_mul [(Nat.choose n i : ℝ), (Nat.choose n (n / 2) : ℝ)] + gcongr ?_ * ?_ + swap + · norm_cast + exact i.choose_le_middle n + specialize hgrowth (n - i) (by simp only [tsub_le_self]) x + refine le_trans (mul_le_mul_of_nonneg_left hgrowth (by positivity)) ?_ + move_mul [C] + gcongr ?_ * C + rw [Finset.mem_range_succ_iff] at hi + change i ≤ (l + k, n).snd at hi + refine le_trans ?_ (one_add_le_sup_seminorm_apply le_rfl hi f x) + rw [pow_add] + move_mul [(1 + ‖x‖) ^ l] + gcongr + simp end Multiplication @@ -824,68 +822,67 @@ variable [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] /-- Composition with a function on the right is a continuous linear map on Schwartz space provided that the function is temperate and growths polynomially near infinity. -/ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) - (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) : 𝓢(E, F) →L[𝕜] 𝓢(D, F) := - mkCLM (fun f x => f (g x)) + (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) : 𝓢(E, F) →L[𝕜] 𝓢(D, F) := by + refine mkCLM (fun f x => f (g x)) (fun _ _ _ => by simp only [add_left_inj, Pi.add_apply, eq_self_iff_true]) (fun _ _ _ => rfl) - (fun f => f.smooth'.comp hg.1) - (by - rintro ⟨k, n⟩ - rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ - rcases hg_upper with ⟨kg, Cg, hg_upper'⟩ - have hCg : 1 ≤ 1 + Cg := by - refine le_add_of_nonneg_right ?_ - specialize hg_upper' 0 - rw [norm_zero] at hg_upper' - exact nonneg_of_mul_nonneg_left hg_upper' (by positivity) - let k' := kg * (k + l * n) - use Finset.Iic (k', n), (1 + Cg) ^ (k + l * n) * ((C + 1) ^ n * n ! * 2 ^ k'), by positivity - intro f x - let seminorm_f := ((Finset.Iic (k', n)).sup (schwartzSeminormFamily 𝕜 _ _)) f - have hg_upper'' : (1 + ‖x‖) ^ (k + l * n) ≤ (1 + Cg) ^ (k + l * n) * (1 + ‖g x‖) ^ k' := by - rw [pow_mul, ← mul_pow] - gcongr - rw [add_mul] - refine add_le_add ?_ (hg_upper' x) - nth_rw 1 [← one_mul (1 : ℝ)] - gcongr - apply one_le_pow₀ - simp only [le_add_iff_nonneg_right, norm_nonneg] - have hbound : - ∀ i, i ≤ n → ‖iteratedFDeriv ℝ i f (g x)‖ ≤ 2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k' := by - intro i hi - have hpos : 0 < (1 + ‖g x‖) ^ k' := by positivity - rw [le_div_iff₀' hpos] - change i ≤ (k', n).snd at hi - exact one_add_le_sup_seminorm_apply le_rfl hi _ _ - have hgrowth' : ∀ N : ℕ, 1 ≤ N → N ≤ n → - ‖iteratedFDeriv ℝ N g x‖ ≤ ((C + 1) * (1 + ‖x‖) ^ l) ^ N := by - intro N hN₁ hN₂ - refine (hgrowth N hN₂ x).trans ?_ - rw [mul_pow] - have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' - gcongr - · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') - · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' - simp only [le_add_iff_nonneg_right, norm_nonneg] - have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' - have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := - pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ - refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ - have rearrange : - (1 + ‖x‖) ^ k * - (n ! * (2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k') * ((C + 1) * (1 + ‖x‖) ^ l) ^ n) = - (1 + ‖x‖) ^ (k + l * n) / (1 + ‖g x‖) ^ k' * - ((C + 1) ^ n * n ! * 2 ^ k' * seminorm_f) := by - rw [mul_pow, pow_add, ← pow_mul] - ring - rw [rearrange] - have hgxk' : 0 < (1 + ‖g x‖) ^ k' := by positivity - rw [← div_le_iff₀ hgxk'] at hg_upper'' - have hpos : (0 : ℝ) ≤ (C + 1) ^ n * n ! * 2 ^ k' * seminorm_f := by - have : 0 ≤ seminorm_f := apply_nonneg _ _ - positivity - refine le_trans (mul_le_mul_of_nonneg_right hg_upper'' hpos) ?_ - rw [← mul_assoc]) + (fun f => f.smooth'.comp hg.1) ?_ + rintro ⟨k, n⟩ + rcases hg.norm_iteratedFDeriv_le_uniform_aux n with ⟨l, C, hC, hgrowth⟩ + rcases hg_upper with ⟨kg, Cg, hg_upper'⟩ + have hCg : 1 ≤ 1 + Cg := by + refine le_add_of_nonneg_right ?_ + specialize hg_upper' 0 + rw [norm_zero] at hg_upper' + exact nonneg_of_mul_nonneg_left hg_upper' (by positivity) + let k' := kg * (k + l * n) + use Finset.Iic (k', n), (1 + Cg) ^ (k + l * n) * ((C + 1) ^ n * n ! * 2 ^ k'), by positivity + intro f x + let seminorm_f := ((Finset.Iic (k', n)).sup (schwartzSeminormFamily 𝕜 _ _)) f + have hg_upper'' : (1 + ‖x‖) ^ (k + l * n) ≤ (1 + Cg) ^ (k + l * n) * (1 + ‖g x‖) ^ k' := by + rw [pow_mul, ← mul_pow] + gcongr + rw [add_mul] + refine add_le_add ?_ (hg_upper' x) + nth_rw 1 [← one_mul (1 : ℝ)] + gcongr + apply one_le_pow₀ + simp only [le_add_iff_nonneg_right, norm_nonneg] + have hbound : + ∀ i, i ≤ n → ‖iteratedFDeriv ℝ i f (g x)‖ ≤ 2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k' := by + intro i hi + have hpos : 0 < (1 + ‖g x‖) ^ k' := by positivity + rw [le_div_iff₀' hpos] + change i ≤ (k', n).snd at hi + exact one_add_le_sup_seminorm_apply le_rfl hi _ _ + have hgrowth' : ∀ N : ℕ, 1 ≤ N → N ≤ n → + ‖iteratedFDeriv ℝ N g x‖ ≤ ((C + 1) * (1 + ‖x‖) ^ l) ^ N := by + intro N hN₁ hN₂ + refine (hgrowth N hN₂ x).trans ?_ + rw [mul_pow] + have hN₁' := (lt_of_lt_of_le zero_lt_one hN₁).ne' + gcongr + · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') + · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' + simp only [le_add_iff_nonneg_right, norm_nonneg] + have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' + have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := + pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ + refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ + have rearrange : + (1 + ‖x‖) ^ k * + (n ! * (2 ^ k' * seminorm_f / (1 + ‖g x‖) ^ k') * ((C + 1) * (1 + ‖x‖) ^ l) ^ n) = + (1 + ‖x‖) ^ (k + l * n) / (1 + ‖g x‖) ^ k' * + ((C + 1) ^ n * n ! * 2 ^ k' * seminorm_f) := by + rw [mul_pow, pow_add, ← pow_mul] + ring + rw [rearrange] + have hgxk' : 0 < (1 + ‖g x‖) ^ k' := by positivity + rw [← div_le_iff₀ hgxk'] at hg_upper'' + have hpos : (0 : ℝ) ≤ (C + 1) ^ n * n ! * 2 ^ k' * seminorm_f := by + have : 0 ≤ seminorm_f := apply_nonneg _ _ + positivity + refine le_trans (mul_le_mul_of_nonneg_right hg_upper'' hpos) ?_ + rw [← mul_assoc] @[simp] lemma compCLM_apply {g : D → E} (hg : g.HasTemperateGrowth) (hg_upper : ∃ (k : ℕ) (C : ℝ), ∀ x, ‖x‖ ≤ C * (1 + ‖g x‖) ^ k) (f : 𝓢(E, F)) : @@ -1070,25 +1067,22 @@ lemma integrable (f : 𝓢(D, V)) : Integrable f μ := variable (𝕜 μ) in /-- The integral as a continuous linear map from Schwartz space to the codomain. -/ -def integralCLM : 𝓢(D, V) →L[𝕜] V := - mkCLMtoNormedSpace (∫ x, · x ∂μ) - (fun f g ↦ by - exact integral_add f.integrable g.integrable) - (integral_smul · ·) - (by - rcases hμ.exists_integrable with ⟨n, h⟩ - let m := (n, 0) - use Finset.Iic m, 2 ^ n * ∫ x : D, (1 + ‖x‖) ^ (- (n : ℝ)) ∂μ - refine ⟨by positivity, fun f ↦ (norm_integral_le_integral_norm f).trans ?_⟩ - have h' : ∀ x, ‖f x‖ ≤ (1 + ‖x‖) ^ (-(n : ℝ)) * - (2 ^ n * ((Finset.Iic m).sup (fun m' => SchwartzMap.seminorm 𝕜 m'.1 m'.2) f)) := by - intro x - rw [rpow_neg (by positivity), ← div_eq_inv_mul, le_div_iff₀' (by positivity), rpow_natCast] - simpa using one_add_le_sup_seminorm_apply (m := m) (k := n) (n := 0) le_rfl le_rfl f x - apply (integral_mono (by simpa using f.integrable_pow_mul μ 0) _ h').trans - · rw [integral_mul_right, ← mul_assoc, mul_comm (2 ^ n)] - rfl - apply h.mul_const) +def integralCLM : 𝓢(D, V) →L[𝕜] V := by + refine mkCLMtoNormedSpace (∫ x, · x ∂μ) + (fun f g ↦ integral_add f.integrable g.integrable) (integral_smul · ·) ?_ + rcases hμ.exists_integrable with ⟨n, h⟩ + let m := (n, 0) + use Finset.Iic m, 2 ^ n * ∫ x : D, (1 + ‖x‖) ^ (- (n : ℝ)) ∂μ + refine ⟨by positivity, fun f ↦ (norm_integral_le_integral_norm f).trans ?_⟩ + have h' : ∀ x, ‖f x‖ ≤ (1 + ‖x‖) ^ (-(n : ℝ)) * + (2 ^ n * ((Finset.Iic m).sup (fun m' => SchwartzMap.seminorm 𝕜 m'.1 m'.2) f)) := by + intro x + rw [rpow_neg (by positivity), ← div_eq_inv_mul, le_div_iff₀' (by positivity), rpow_natCast] + simpa using one_add_le_sup_seminorm_apply (m := m) (k := n) (n := 0) le_rfl le_rfl f x + apply (integral_mono (by simpa using f.integrable_pow_mul μ 0) _ h').trans + · rw [integral_mul_right, ← mul_assoc, mul_comm (2 ^ n)] + rfl + apply h.mul_const variable (𝕜) in @[simp] From 2bb0410843932d03ff2c12cd9bf2e996f514a36b Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Fri, 22 Nov 2024 05:55:44 +0000 Subject: [PATCH 025/250] chore(discover-lean-pr-testing): fix multiline output to GITHUB_OUTPUT (#19355) --- .github/workflows/discover-lean-pr-testing.yml | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 69e46265bbbc0..2bee99c8e421d 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -4,8 +4,6 @@ on: push: branches: - nightly-testing - paths: - - lean-toolchain jobs: discover-lean-pr-testing: @@ -69,7 +67,7 @@ jobs: # Output the PRs echo "$PRS" - echo "prs=$PRS" >> "$GITHUB_OUTPUT" + printf "prs<> "$GITHUB_OUTPUT" - name: Use merged PRs information id: find-branches From af1911940422ded96480cde9daed0a5de7a435d3 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 22 Nov 2024 08:14:07 +0000 Subject: [PATCH 026/250] chore: add shortcut instances after LinearOrder Nat (#19346) This prevents Lean from going through the Lattice structure, which requires additional imports. --- Mathlib/Algebra/Prime/Lemmas.lean | 2 +- Mathlib/CategoryTheory/Functor/OfSequence.lean | 1 - Mathlib/Data/Int/GCD.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 7 +++++++ Mathlib/Data/Nat/Log.lean | 6 +++--- Mathlib/Data/Nat/Prime/Defs.lean | 2 +- Mathlib/Data/Nat/Size.lean | 1 - 7 files changed, 13 insertions(+), 8 deletions(-) diff --git a/Mathlib/Algebra/Prime/Lemmas.lean b/Mathlib/Algebra/Prime/Lemmas.lean index 96ee461f3831f..0211c3bbc8c5e 100644 --- a/Mathlib/Algebra/Prime/Lemmas.lean +++ b/Mathlib/Algebra/Prime/Lemmas.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.Group.Even import Mathlib.Algebra.Group.Units.Equiv import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Prime.Defs -import Mathlib.Order.Lattice +import Mathlib.Order.Monotone.Basic /-! # Associated, prime, and irreducible elements. diff --git a/Mathlib/CategoryTheory/Functor/OfSequence.lean b/Mathlib/CategoryTheory/Functor/OfSequence.lean index ba70972cf1661..c1f5788cec33b 100644 --- a/Mathlib/CategoryTheory/Functor/OfSequence.lean +++ b/Mathlib/CategoryTheory/Functor/OfSequence.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.Algebra.Order.Group.Nat import Mathlib.CategoryTheory.Category.Preorder import Mathlib.CategoryTheory.EqToHom diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index eccc02d840c40..e6b7b6ec9f4e8 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -7,8 +7,8 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.GroupWithZero.Semiconj import Mathlib.Algebra.Group.Commute.Units import Mathlib.Data.Nat.GCD.Basic -import Mathlib.Order.Lattice import Mathlib.Data.Set.Operations +import Mathlib.Order.Basic import Mathlib.Order.Bounds.Defs /-! diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 9063d971434ff..2327e2c3bb37a 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -72,6 +72,13 @@ instance instLinearOrder : LinearOrder ℕ where decidableLE := inferInstance decidableEq := inferInstance +-- Shortcut instances +instance : Preorder ℕ := inferInstance +instance : PartialOrder ℕ := inferInstance +instance : Min ℕ := inferInstance +instance : Max ℕ := inferInstance +instance : Ord ℕ := inferInstance + instance instNontrivial : Nontrivial ℕ := ⟨⟨0, 1, Nat.zero_ne_one⟩⟩ @[simp] theorem default_eq_zero : default = 0 := rfl diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index fc11e395a0c41..1d16748dda19e 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -3,11 +3,11 @@ Copyright (c) 2020 Simon Hudon. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Simon Hudon, Yaël Dillies -/ +import Mathlib.Order.Interval.Set.Defs +import Mathlib.Order.Monotone.Basic import Mathlib.Tactic.Bound.Attribute -import Mathlib.Tactic.Monotonicity.Attr -import Mathlib.Order.Lattice import Mathlib.Tactic.Contrapose -import Mathlib.Order.Interval.Set.Defs +import Mathlib.Tactic.Monotonicity.Attr /-! # Natural number logarithms diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index 30b5a96a46888..c978c30544721 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -6,7 +6,7 @@ Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro import Batteries.Data.Nat.Gcd import Mathlib.Algebra.Prime.Defs import Mathlib.Algebra.Ring.Nat -import Mathlib.Order.Lattice +import Mathlib.Order.Basic /-! # Prime numbers diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index fec3892450d06..fd9311b427815 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.Nat.Bits -import Mathlib.Order.Lattice /-! Lemmas about `size`. -/ From 29729770e05e42140e678ced52f740a36019e3a7 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Fri, 22 Nov 2024 08:33:47 +0000 Subject: [PATCH 027/250] chore: bump to nightly-2024-11-22 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index f589f3bf8fa9a..ae2a1d4650fc0 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-21 +leanprover/lean4:nightly-2024-11-22 From 8e3e89e9b3c5b0a719dea497f46946c860381031 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 22 Nov 2024 08:58:22 +0000 Subject: [PATCH 028/250] chore: split Algebra.Group.TypeTags (#19351) --- Mathlib.lean | 4 +- Mathlib/Algebra/Category/Grp/FiniteGrp.lean | 1 + Mathlib/Algebra/FreeMonoid/Count.lean | 1 + Mathlib/Algebra/Group/Action/TypeTags.lean | 2 +- Mathlib/Algebra/Group/Equiv/TypeTags.lean | 2 +- Mathlib/Algebra/Group/Even.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Map.lean | 1 + .../Group/Subsemigroup/Operations.lean | 2 +- .../{TypeTags.lean => TypeTags/Basic.lean} | 124 ++---------------- Mathlib/Algebra/Group/TypeTags/Finite.lean | 28 ++++ Mathlib/Algebra/Group/TypeTags/Hom.lean | 123 +++++++++++++++++ .../Order/Monoid/Unbundled/TypeTags.lean | 2 +- Mathlib/Analysis/Normed/Field/Lemmas.lean | 1 + Mathlib/Data/Finsupp/Defs.lean | 1 + Mathlib/Data/Fintype/Basic.lean | 1 + Mathlib/Data/Nat/Cast/Basic.lean | 1 + Mathlib/Deprecated/Group.lean | 2 +- Mathlib/GroupTheory/FiniteAbelian/Basic.lean | 1 + 18 files changed, 176 insertions(+), 123 deletions(-) rename Mathlib/Algebra/Group/{TypeTags.lean => TypeTags/Basic.lean} (80%) create mode 100644 Mathlib/Algebra/Group/TypeTags/Finite.lean create mode 100644 Mathlib/Algebra/Group/TypeTags/Hom.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5e0aee1763797..55552ad2a0f9b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -321,7 +321,9 @@ import Mathlib.Algebra.Group.Subsemigroup.Membership import Mathlib.Algebra.Group.Subsemigroup.Operations import Mathlib.Algebra.Group.Support import Mathlib.Algebra.Group.Translate -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic +import Mathlib.Algebra.Group.TypeTags.Finite +import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Group.ULift import Mathlib.Algebra.Group.UniqueProds.Basic import Mathlib.Algebra.Group.UniqueProds.VectorSpace diff --git a/Mathlib/Algebra/Category/Grp/FiniteGrp.lean b/Mathlib/Algebra/Category/Grp/FiniteGrp.lean index b1f4d7ead61ee..ac01e3ac8ce3c 100644 --- a/Mathlib/Algebra/Category/Grp/FiniteGrp.lean +++ b/Mathlib/Algebra/Category/Grp/FiniteGrp.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 Jujian Zhang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jujian Zhang, Nailin Guan, Yuyang Zhao -/ +import Mathlib.Data.Finite.Defs import Mathlib.Algebra.Category.Grp.Basic /-! diff --git a/Mathlib/Algebra/FreeMonoid/Count.lean b/Mathlib/Algebra/FreeMonoid/Count.lean index df4dcfd27f672..53b2c0e4c51c3 100644 --- a/Mathlib/Algebra/FreeMonoid/Count.lean +++ b/Mathlib/Algebra/FreeMonoid/Count.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.FreeMonoid.Basic +import Mathlib.Algebra.Group.TypeTags.Hom /-! # `List.count` as a bundled homomorphism diff --git a/Mathlib/Algebra/Group/Action/TypeTags.lean b/Mathlib/Algebra/Group/Action/TypeTags.lean index 0fa009b36d114..c6ceefe332560 100644 --- a/Mathlib/Algebra/Group/Action/TypeTags.lean +++ b/Mathlib/Algebra/Group/Action/TypeTags.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.End -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Additive and Multiplicative for group actions diff --git a/Mathlib/Algebra/Group/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean index 78c085f074666..f8e14c0ceb128 100644 --- a/Mathlib/Algebra/Group/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -5,7 +5,7 @@ Authors: Johannes Hölzl, Callum Sutton, Yury Kudryashov -/ import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Prod -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Additive and multiplicative equivalences associated to `Multiplicative` and `Additive`. diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index 5bd24647c2f65..febd4c5c50c99 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ import Mathlib.Algebra.Group.Opposite -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Squares and even elements diff --git a/Mathlib/Algebra/Group/Subgroup/Map.lean b/Mathlib/Algebra/Group/Subgroup/Map.lean index 84450e745384e..e1219dcd59203 100644 --- a/Mathlib/Algebra/Group/Subgroup/Map.lean +++ b/Mathlib/Algebra/Group/Subgroup/Map.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kexing Ying -/ import Mathlib.Algebra.Group.Subgroup.Lattice +import Mathlib.Algebra.Group.TypeTags.Hom /-! # `map` and `comap` for subgroups diff --git a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean index 8fc45916faae0..2c782614c04cf 100644 --- a/Mathlib/Algebra/Group/Subsemigroup/Operations.lean +++ b/Mathlib/Algebra/Group/Subsemigroup/Operations.lean @@ -6,7 +6,7 @@ Amelia Livingston, Yury Kudryashov, Yakov Pechersky, Jireh Loreaux -/ import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Subsemigroup.Basic -import Mathlib.Algebra.Group.TypeTags +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Operations on `Subsemigroup`s diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags/Basic.lean similarity index 80% rename from Mathlib/Algebra/Group/TypeTags.lean rename to Mathlib/Algebra/Group/TypeTags/Basic.lean index f6e2f9c2cc9e0..c506dc1cf3541 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags/Basic.lean @@ -3,8 +3,12 @@ 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.Group.Hom.Defs -import Mathlib.Data.Finite.Defs +import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Data.FunLike.Basic +import Mathlib.Logic.Function.Iterate +import Mathlib.Logic.Equiv.Defs +import Mathlib.Tactic.Set +import Mathlib.Util.AssertExists import Mathlib.Logic.Nontrivial.Basic /-! @@ -32,6 +36,8 @@ This file is similar to `Order.Synonym`. assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered +assert_not_exists MonoidHom +assert_not_exists Finite universe u v @@ -136,16 +142,6 @@ instance [Inhabited α] : Inhabited (Multiplicative α) := instance [Unique α] : Unique (Additive α) := toMul.unique instance [Unique α] : Unique (Multiplicative α) := toAdd.unique -instance [Finite α] : Finite (Additive α) := - Finite.of_equiv α (by rfl) - -instance [Finite α] : Finite (Multiplicative α) := - Finite.of_equiv α (by rfl) - -instance [h : Infinite α] : Infinite (Additive α) := h - -instance [h : Infinite α] : Infinite (Multiplicative α) := h - instance [h : DecidableEq α] : DecidableEq (Multiplicative α) := h instance [h : DecidableEq α] : DecidableEq (Additive α) := h @@ -431,110 +427,6 @@ instance Additive.addCommGroup [CommGroup α] : AddCommGroup (Additive α) := instance Multiplicative.commGroup [AddCommGroup α] : CommGroup (Multiplicative α) := { Multiplicative.group, Multiplicative.commMonoid with } -/-- Reinterpret `α →+ β` as `Multiplicative α →* Multiplicative β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : - (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where - toFun f := { - toFun := fun a => ofAdd (f (toAdd a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => toAdd (f (ofAdd a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : α →+ β) : - ⇑(toMultiplicative f) = ofAdd ∘ f ∘ toAdd := rfl - -/-- Reinterpret `α →* β` as `Additive α →+ Additive β`. -/ -@[simps] -def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : - (α →* β) ≃ (Additive α →+ Additive β) where - toFun f := { - toFun := fun a => ofMul (f (toMul a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - invFun f := { - toFun := fun a => toMul (f (ofMul a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma MonoidHom.coe_toMultiplicative [MulOneClass α] [MulOneClass β] (f : α →* β) : - ⇑(toAdditive f) = ofMul ∘ f ∘ toMul := rfl - -/-- Reinterpret `Additive α →+ β` as `α →* Multiplicative β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : - (Additive α →+ β) ≃ (α →* Multiplicative β) where - toFun f := { - toFun := fun a => ofAdd (f (ofMul a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => toAdd (f (toMul a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative' [MulOneClass α] [AddZeroClass β] (f : Additive α →+ β) : - ⇑(toMultiplicative' f) = ofAdd ∘ f ∘ ofMul := rfl - -/-- Reinterpret `α →* Multiplicative β` as `Additive α →+ β`. -/ -@[simps!] -def MonoidHom.toAdditive' [MulOneClass α] [AddZeroClass β] : - (α →* Multiplicative β) ≃ (Additive α →+ β) := - AddMonoidHom.toMultiplicative'.symm - -@[simp, norm_cast] -lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* Multiplicative β) : - ⇑(toAdditive' f) = toAdd ∘ f ∘ toMul := rfl - -/-- Reinterpret `α →+ Additive β` as `Multiplicative α →* β`. -/ -@[simps] -def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : - (α →+ Additive β) ≃ (Multiplicative α →* β) where - toFun f := { - toFun := fun a => toMul (f (toAdd a)) - map_mul' := f.map_add - map_one' := f.map_zero - } - invFun f := { - toFun := fun a => ofMul (f (ofAdd a)) - map_add' := f.map_mul - map_zero' := f.map_one - } - left_inv _ := rfl - right_inv _ := rfl - -@[simp, norm_cast] -lemma AddMonoidHom.coe_toMultiplicative'' [AddZeroClass α] [MulOneClass β] (f : α →+ Additive β) : - ⇑(toMultiplicative'' f) = toMul ∘ f ∘ toAdd := rfl - -/-- Reinterpret `Multiplicative α →* β` as `α →+ Additive β`. -/ -@[simps!] -def MonoidHom.toAdditive'' [AddZeroClass α] [MulOneClass β] : - (Multiplicative α →* β) ≃ (α →+ Additive β) := - AddMonoidHom.toMultiplicative''.symm - -@[simp, norm_cast] -lemma MonoidHom.coe_toAdditive'' [AddZeroClass α] [MulOneClass β] (f : Multiplicative α →* β) : - ⇑(toAdditive'' f) = ofMul ∘ f ∘ ofAdd := rfl - /-- If `α` has some multiplicative structure and coerces to a function, then `Additive α` should also coerce to the same function. diff --git a/Mathlib/Algebra/Group/TypeTags/Finite.lean b/Mathlib/Algebra/Group/TypeTags/Finite.lean new file mode 100644 index 0000000000000..e3885696b373b --- /dev/null +++ b/Mathlib/Algebra/Group/TypeTags/Finite.lean @@ -0,0 +1,28 @@ +/- +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.Group.TypeTags.Basic +import Mathlib.Data.Finite.Defs + +/-! +# `Finite` and `Infinite` are preserved by `Additive` and `Multiplicative`. +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +universe u v + +variable {α : Type u} + +instance [Finite α] : Finite (Additive α) := + Finite.of_equiv α (by rfl) + +instance [Finite α] : Finite (Multiplicative α) := + Finite.of_equiv α (by rfl) + +instance [h : Infinite α] : Infinite (Additive α) := h + +instance [h : Infinite α] : Infinite (Multiplicative α) := h diff --git a/Mathlib/Algebra/Group/TypeTags/Hom.lean b/Mathlib/Algebra/Group/TypeTags/Hom.lean new file mode 100644 index 0000000000000..749c54aa7a8b7 --- /dev/null +++ b/Mathlib/Algebra/Group/TypeTags/Hom.lean @@ -0,0 +1,123 @@ +/- +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.Group.Hom.Defs +import Mathlib.Algebra.Group.TypeTags.Basic + +/-! +# Transport algebra morphisms between additive and multiplicative types. +-/ + + +universe u v + +variable {α : Type u} {β : Type v} + +open Additive (ofMul toMul) +open Multiplicative (ofAdd toAdd) + +/-- Reinterpret `α →+ β` as `Multiplicative α →* Multiplicative β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : + (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where + toFun f := { + toFun := fun a => ofAdd (f (toAdd a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => toAdd (f (ofAdd a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : α →+ β) : + ⇑(toMultiplicative f) = ofAdd ∘ f ∘ toAdd := rfl + +/-- Reinterpret `α →* β` as `Additive α →+ Additive β`. -/ +@[simps] +def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : + (α →* β) ≃ (Additive α →+ Additive β) where + toFun f := { + toFun := fun a => ofMul (f (toMul a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + invFun f := { + toFun := fun a => toMul (f (ofMul a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma MonoidHom.coe_toMultiplicative [MulOneClass α] [MulOneClass β] (f : α →* β) : + ⇑(toAdditive f) = ofMul ∘ f ∘ toMul := rfl + +/-- Reinterpret `Additive α →+ β` as `α →* Multiplicative β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : + (Additive α →+ β) ≃ (α →* Multiplicative β) where + toFun f := { + toFun := fun a => ofAdd (f (ofMul a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => toAdd (f (toMul a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative' [MulOneClass α] [AddZeroClass β] (f : Additive α →+ β) : + ⇑(toMultiplicative' f) = ofAdd ∘ f ∘ ofMul := rfl + +/-- Reinterpret `α →* Multiplicative β` as `Additive α →+ β`. -/ +@[simps!] +def MonoidHom.toAdditive' [MulOneClass α] [AddZeroClass β] : + (α →* Multiplicative β) ≃ (Additive α →+ β) := + AddMonoidHom.toMultiplicative'.symm + +@[simp, norm_cast] +lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* Multiplicative β) : + ⇑(toAdditive' f) = toAdd ∘ f ∘ toMul := rfl + +/-- Reinterpret `α →+ Additive β` as `Multiplicative α →* β`. -/ +@[simps] +def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : + (α →+ Additive β) ≃ (Multiplicative α →* β) where + toFun f := { + toFun := fun a => toMul (f (toAdd a)) + map_mul' := f.map_add + map_one' := f.map_zero + } + invFun f := { + toFun := fun a => ofMul (f (ofAdd a)) + map_add' := f.map_mul + map_zero' := f.map_one + } + left_inv _ := rfl + right_inv _ := rfl + +@[simp, norm_cast] +lemma AddMonoidHom.coe_toMultiplicative'' [AddZeroClass α] [MulOneClass β] (f : α →+ Additive β) : + ⇑(toMultiplicative'' f) = toMul ∘ f ∘ toAdd := rfl + +/-- Reinterpret `Multiplicative α →* β` as `α →+ Additive β`. -/ +@[simps!] +def MonoidHom.toAdditive'' [AddZeroClass α] [MulOneClass β] : + (Multiplicative α →* β) ≃ (α →+ Additive β) := + AddMonoidHom.toMultiplicative''.symm + +@[simp, norm_cast] +lemma MonoidHom.coe_toAdditive'' [AddZeroClass α] [MulOneClass β] (f : Multiplicative α →* β) : + ⇑(toAdditive'' f) = ofMul ∘ f ∘ ofAdd := rfl diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index 1647d43c6a434..a975f6b4f3384 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -3,9 +3,9 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Order.BoundedOrder +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index e22f72874193c..88affc76dbe77 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -5,6 +5,7 @@ Authors: Patrick Massot, Johannes Hölzl -/ import Mathlib.Algebra.Group.AddChar +import Mathlib.Algebra.Group.TypeTags.Finite import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Group.Bounded diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 7ac40e922fd7b..be9894c31daf0 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Group.Indicator import Mathlib.Algebra.Group.Submonoid.Basic import Mathlib.Data.Finset.Max import Mathlib.Data.Set.Finite.Basic +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Type of functions with finite support diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 3bda18e430ecb..d77529bf9ac3e 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro -/ import Mathlib.Data.Finset.Image import Mathlib.Data.List.FinRange +import Mathlib.Data.Finite.Defs /-! # Finite types diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index d4f2aa9e4fd08..b8cda137de314 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -6,6 +6,7 @@ Authors: Mario Carneiro import Mathlib.Algebra.Divisibility.Basic import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Nat +import Mathlib.Algebra.Group.TypeTags.Hom /-! # Cast of natural numbers (additional theorems) diff --git a/Mathlib/Deprecated/Group.lean b/Mathlib/Deprecated/Group.lean index 539782f49e725..66bc94fe1317e 100644 --- a/Mathlib/Deprecated/Group.lean +++ b/Mathlib/Deprecated/Group.lean @@ -4,9 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Equiv.Basic -import Mathlib.Algebra.Group.TypeTags import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.Ring.Hom.Defs +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Unbundled monoid and group homomorphisms diff --git a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean index def95dc958b3b..38fa544799b25 100644 --- a/Mathlib/GroupTheory/FiniteAbelian/Basic.lean +++ b/Mathlib/GroupTheory/FiniteAbelian/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Pierre-Alexandre Bazin -/ import Mathlib.Algebra.Module.PID +import Mathlib.Algebra.Group.TypeTags.Finite import Mathlib.Data.ZMod.Quotient /-! From 944e3576a3e6deb8a0b61ace4c5df6b80a47435f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Fri, 22 Nov 2024 09:37:25 +0000 Subject: [PATCH 029/250] chore: move def MonoidHom.inverse to earlier defs file (#19348) No import effect, just moving defs to be with other defs. --- Mathlib/Algebra/Group/Equiv/Basic.lean | 31 -------------------------- Mathlib/Algebra/Group/Hom/Defs.lean | 31 ++++++++++++++++++++++++++ MathlibTest/symm.lean | 9 -------- 3 files changed, 31 insertions(+), 40 deletions(-) diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 81aff3106d93e..203ef93e84ead 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -45,37 +45,6 @@ theorem map_ne_one_iff {f : F} {x : M} : end EmbeddingLike -/-- Makes a `OneHom` inverse from the bijective inverse of a `OneHom` -/ -@[to_additive (attr := simps) - "Make a `ZeroHom` inverse from the bijective inverse of a `ZeroHom`"] -def OneHom.inverse [One M] [One N] - (f : OneHom M N) (g : N → M) - (h₁ : Function.LeftInverse g f) : - OneHom N M := - { toFun := g, - map_one' := by rw [← f.map_one, h₁] } - -/-- Makes a multiplicative inverse from a bijection which preserves multiplication. -/ -@[to_additive (attr := simps) - "Makes an additive inverse from a bijection which preserves addition."] -def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) - (h₁ : Function.LeftInverse g f) - (h₂ : Function.RightInverse g f) : N →ₙ* M where - toFun := g - map_mul' x y := - calc - g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] - _ = g (f (g x * g y)) := by rw [f.map_mul] - _ = g x * g y := h₁ _ - -/-- The inverse of a bijective `MonoidHom` is a `MonoidHom`. -/ -@[to_additive (attr := simps) - "The inverse of a bijective `AddMonoidHom` is an `AddMonoidHom`."] -def MonoidHom.inverse {A B : Type*} [Monoid A] [Monoid B] (f : A →* B) (g : B → A) - (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : B →* A := - { (f : OneHom A B).inverse g h₁, - (f : A →ₙ* B).inverse g h₁ h₂ with toFun := g } - /-- `AddEquiv α β` is the type of an equiv `α ≃ β` which preserves addition. -/ structure AddEquiv (A B : Type*) [Add A] [Add B] extends A ≃ B, AddHom A B diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index 14f51627f6d9d..f2540ca0f00ab 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -843,6 +843,37 @@ protected theorem MonoidHom.map_zpow' [DivInvMonoid M] [DivInvMonoid N] (f : M (hf : ∀ x, f x⁻¹ = (f x)⁻¹) (a : M) (n : ℤ) : f (a ^ n) = f a ^ n := map_zpow' f hf a n +/-- Makes a `OneHom` inverse from the bijective inverse of a `OneHom` -/ +@[to_additive (attr := simps) + "Make a `ZeroHom` inverse from the bijective inverse of a `ZeroHom`"] +def OneHom.inverse [One M] [One N] + (f : OneHom M N) (g : N → M) + (h₁ : Function.LeftInverse g f) : + OneHom N M := + { toFun := g, + map_one' := by rw [← f.map_one, h₁] } + +/-- Makes a multiplicative inverse from a bijection which preserves multiplication. -/ +@[to_additive (attr := simps) + "Makes an additive inverse from a bijection which preserves addition."] +def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) + (h₁ : Function.LeftInverse g f) + (h₂ : Function.RightInverse g f) : N →ₙ* M where + toFun := g + map_mul' x y := + calc + g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] + _ = g (f (g x * g y)) := by rw [f.map_mul] + _ = g x * g y := h₁ _ + +/-- The inverse of a bijective `MonoidHom` is a `MonoidHom`. -/ +@[to_additive (attr := simps) + "The inverse of a bijective `AddMonoidHom` is an `AddMonoidHom`."] +def MonoidHom.inverse {A B : Type*} [Monoid A] [Monoid B] (f : A →* B) (g : B → A) + (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : B →* A := + { (f : OneHom A B).inverse g h₁, + (f : A →ₙ* B).inverse g h₁ h₂ with toFun := g } + section End namespace Monoid diff --git a/MathlibTest/symm.lean b/MathlibTest/symm.lean index 7d88a90868362..ae72925a5b7a4 100644 --- a/MathlibTest/symm.lean +++ b/MathlibTest/symm.lean @@ -23,15 +23,6 @@ example (a b c : Nat) (ab : a = b) (bc : b = c) : c = a := by symm_saturate apply Eq.trans <;> assumption -def MulHom.inverse [Mul M] [Mul N] (f : M →ₙ* N) (g : N → M) (h₁ : Function.LeftInverse g f) - (h₂ : Function.RightInverse g f) : N →ₙ* M where - toFun := g - map_mul' x y := - calc - g (x * y) = g (f (g x) * f (g y)) := by rw [h₂ x, h₂ y] - _ = g (f (g x * g y)) := by rw [f.map_mul] - _ = g x * g y := h₁ _ - structure MulEquiv (M N : Type u) [Mul M] [Mul N] extends M ≃ N, M →ₙ* N /-- From 1a2513aa6adbca25ceb22c79fea826dd0acb93f3 Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Fri, 22 Nov 2024 10:20:46 +0000 Subject: [PATCH 030/250] fix: move `Y` notation into a deeper scope (#19292) Otherwise this aggressively claims the `Y` notation in places where the intent may only be to access `Polynomial.X`. An alternative would be to make `Y` an `abbrev`, which would cause it to participate in regular name resolution, allowing local variables to be called `Y`. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Aggressive.20notation.2C.20Y/near/483466695) --- Mathlib/Algebra/Polynomial/Bivariate.lean | 6 ++++-- Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean | 1 + .../EllipticCurve/DivisionPolynomial/Basic.lean | 1 + Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean | 1 + 4 files changed, 7 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Bivariate.lean b/Mathlib/Algebra/Polynomial/Bivariate.lean index a78e8ffa4e2c8..d9feee54b8b5c 100644 --- a/Mathlib/Algebra/Polynomial/Bivariate.lean +++ b/Mathlib/Algebra/Polynomial/Bivariate.lean @@ -17,10 +17,12 @@ the abbreviation `CC` to view a constant in the base ring `R` as a bivariate pol -/ /-- The notation `Y` for `X` in the `Polynomial` scope. -/ -scoped[Polynomial] notation3:max "Y" => Polynomial.X (R := Polynomial _) +scoped[Polynomial.Bivariate] notation3:max "Y" => Polynomial.X (R := Polynomial _) /-- The notation `R[X][Y]` for `R[X][X]` in the `Polynomial` scope. -/ -scoped[Polynomial] notation3:max R "[X][Y]" => Polynomial (Polynomial R) +scoped[Polynomial.Bivariate] notation3:max R "[X][Y]" => Polynomial (Polynomial R) + +open scoped Polynomial.Bivariate namespace Polynomial diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean index 5a23dde317010..231b6e92006e9 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Affine.lean @@ -81,6 +81,7 @@ elliptic curve, rational point, affine coordinates -/ open Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean index 5fb7a921a4ae8..15a2fccfcd2b3 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/DivisionPolynomial/Basic.lean @@ -95,6 +95,7 @@ elliptic curve, division polynomial, torsion point -/ open Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) diff --git a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean index 7020898f0efde..c01a58c13679a 100644 --- a/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean +++ b/Mathlib/AlgebraicGeometry/EllipticCurve/Group.lean @@ -62,6 +62,7 @@ elliptic curve, group law, class group -/ open Ideal nonZeroDivisors Polynomial +open scoped Polynomial.Bivariate local macro "C_simp" : tactic => `(tactic| simp only [map_ofNat, C_0, C_1, C_neg, C_add, C_sub, C_mul, C_pow]) From d511775f5a2ce12b2d54913c5f9ffdc0457bb663 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Fri, 22 Nov 2024 10:20:47 +0000 Subject: [PATCH 031/250] docs(RingTheory/Finiteness/Defs): specify namespace of `Finite` (#19359) Currently, `Module.Finite` and `RingHom.Finite` have docstrings that link to `Finite`, which is confusing. --- Mathlib/RingTheory/Finiteness/Defs.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/RingTheory/Finiteness/Defs.lean b/Mathlib/RingTheory/Finiteness/Defs.lean index 289071d1c7064..a56606fa8c8e9 100644 --- a/Mathlib/RingTheory/Finiteness/Defs.lean +++ b/Mathlib/RingTheory/Finiteness/Defs.lean @@ -100,7 +100,7 @@ section ModuleAndAlgebra variable (R A B M N : Type*) -/-- A module over a semiring is `Finite` if it is finitely generated as a module. -/ +/-- A module over a semiring is `Module.Finite` if it is finitely generated as a module. -/ protected class Module.Finite [Semiring R] [AddCommMonoid M] [Module R M] : Prop where out : (⊤ : Submodule R M).FG @@ -142,7 +142,7 @@ namespace RingHom variable {A B C : Type*} [CommRing A] [CommRing B] [CommRing C] -/-- A ring morphism `A →+* B` is `Finite` if `B` is finitely generated as `A`-module. -/ +/-- A ring morphism `A →+* B` is `RingHom.Finite` if `B` is finitely generated as `A`-module. -/ @[algebraize Module.Finite] def Finite (f : A →+* B) : Prop := letI : Algebra A B := f.toAlgebra From dc8fc4ec255093931e2d5370201349d86d3025e2 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 22 Nov 2024 10:42:59 +0000 Subject: [PATCH 032/250] feat: `Matrix.abs_det_submatrix_equiv_equiv` (#19332) --- Mathlib/Algebra/Order/Ring/Abs.lean | 4 ++++ Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 11 +++++++++++ 2 files changed, 15 insertions(+) diff --git a/Mathlib/Algebra/Order/Ring/Abs.lean b/Mathlib/Algebra/Order/Ring/Abs.lean index 9948d23da4d4b..5552705b81dff 100644 --- a/Mathlib/Algebra/Order/Ring/Abs.lean +++ b/Mathlib/Algebra/Order/Ring/Abs.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.Order.Group.Abs import Mathlib.Algebra.Order.Ring.Basic import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Ring.Divisibility.Basic +import Mathlib.Algebra.Ring.Int.Units import Mathlib.Data.Nat.Cast.Order.Ring /-! @@ -151,6 +152,9 @@ theorem abs_sub_sq (a b : α) : |a - b| * |a - b| = a * a + b * b - (1 + 1) * a simp only [mul_add, add_comm, add_left_comm, mul_comm, sub_eq_add_neg, mul_one, mul_neg, neg_add_rev, neg_neg, add_assoc] +lemma abs_unit_intCast (a : ℤˣ) : |((a : ℤ) : α)| = 1 := by + cases Int.units_eq_one_or a <;> simp_all + end LinearOrderedCommRing section diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 1001a688e6203..2571ccd9ad821 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -222,6 +222,17 @@ theorem det_submatrix_equiv_self (e : n ≃ m) (A : Matrix m m R) : intro i rw [Equiv.permCongr_apply, Equiv.symm_apply_apply, submatrix_apply] +/-- Permuting rows and columns with two equivalences does not change the absolute value of the +determinant. -/ +@[simp] +theorem abs_det_submatrix_equiv_equiv {R : Type*} [LinearOrderedCommRing R] + (e₁ e₂ : n ≃ m) (A : Matrix m m R) : + |(A.submatrix e₁ e₂).det| = |A.det| := by + have hee : e₂ = e₁.trans (e₁.symm.trans e₂) := by ext; simp + rw [hee] + show |((A.submatrix id (e₁.symm.trans e₂)).submatrix e₁ e₁).det| = |A.det| + rw [Matrix.det_submatrix_equiv_self, Matrix.det_permute', abs_mul, abs_unit_intCast, one_mul] + /-- Reindexing both indices along the same equivalence preserves the determinant. For the `simp` version of this lemma, see `det_submatrix_equiv_self`; this one is unsuitable because From 81381ad8339f8f2b075d0846b3d11fd3ecca4e75 Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" Date: Fri, 22 Nov 2024 11:21:14 +0000 Subject: [PATCH 033/250] doc(Algebra/Central/Defs): correct doc (#19363) Correct two typos and a bad syntax in the doc. --- Mathlib/Algebra/Central/Defs.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Central/Defs.lean b/Mathlib/Algebra/Central/Defs.lean index c88a10b8b3aba..363464e22abef 100644 --- a/Mathlib/Algebra/Central/Defs.lean +++ b/Mathlib/Algebra/Central/Defs.lean @@ -19,7 +19,7 @@ is a (not necessarily commutative) `K`-algebra. ## Implementation notes We require the `K`-center of `D` to be smaller than or equal to the smallest subalgebra so that when -we prove something is central, there we don't need to prove `⊥ ≤ center K D` even though this +we prove something is central, we don't need to prove `⊥ ≤ center K D` even though this direction is trivial. ### Central Simple Algebras @@ -34,11 +34,11 @@ but an instance of `[Algebra.IsCentralSimple K D]` would not imply `[IsSimpleRin synthesization orders (`K` cannot be inferred). Thus, to obtain a central simple `K`-algebra `D`, one should use `Algebra.IsCentral K D` and `IsSimpleRing D` separately. -Note that the predicate `Albgera.IsCentral K D` and `IsSimpleRing D` makes sense just for `K` a +Note that the predicate `Algebra.IsCentral K D` and `IsSimpleRing D` makes sense just for `K` a `CommRing` but it doesn't give the right definition for central simple algebra; for a commutative ring base, one should use the theory of Azumaya algebras. In fact ideals of `K` immediately give rise to nontrivial quotients of `D` so there are no central simple algebras in this case according -to our definition, if K is not a field. +to our definition, if `K` is not a field. The theory of central simple algebras really is a theory over fields. Thus to declare a central simple algebra, one should use the following: From 0c44f3226511eec8ec0fbef477cc35921fe0fa2d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Martin=20Dvo=C5=99=C3=A1k?= Date: Fri, 22 Nov 2024 11:41:25 +0000 Subject: [PATCH 034/250] docs(LinearAlgebra/Matrix/Determinant/Basic): det_submatrix_equiv_self (#19365) --- Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean index 2571ccd9ad821..bd291498d65eb 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/Basic.lean @@ -209,7 +209,7 @@ theorem det_permute' (σ : Perm n) (M : Matrix n n R) : (M.submatrix id σ).det = Perm.sign σ * M.det := by rw [← det_transpose, transpose_submatrix, det_permute, det_transpose] -/-- Permuting rows and columns with the same equivalence has no effect. -/ +/-- Permuting rows and columns with the same equivalence does not change the determinant. -/ @[simp] theorem det_submatrix_equiv_self (e : n ≃ m) (A : Matrix m m R) : det (A.submatrix e e) = det A := by From d380d2c97aba4eadfd37855ce5fb64d87caafd3d Mon Sep 17 00:00:00 2001 From: "Filippo A. E. Nuccio" <109365723+Yu-Misaka@users.noreply.github.com> Date: Fri, 22 Nov 2024 13:03:58 +0000 Subject: [PATCH 035/250] feat(FieldTheory.JacobsonNoether) : add proof of the Jacobson-Noether theorem (#16525) In this PR we provide a proof of the Jacobson-Noether theorem following [this blog](https://ysharifi.wordpress.com/2011/09/30/the-jacobson-noether-theorem/) Co-authored by: Filippo A. E. Nuccio @faenuccio Co-authored by: Wanyi He @Blackfeather007 Co-authored by: Huanyu Zheng @Yu-Misaka Co-authored by: Yi Yuan @yuanyi-350 Co-authored by: Weichen Jiao @AlbertJ-314 Co-authored by: Sihan Wu @Old-Turtledove Co-authored by: @FR-vdash-bot Co-authored-by: faenuccio --- Mathlib.lean | 2 + Mathlib/Algebra/Algebra/Basic.lean | 14 +- Mathlib/Algebra/CharP/LinearMaps.lean | 60 +++++++ Mathlib/Algebra/CharP/Subring.lean | 17 +- Mathlib/Algebra/GroupWithZero/Conj.lean | 8 + Mathlib/FieldTheory/JacobsonNoether.lean | 195 +++++++++++++++++++++++ 6 files changed, 290 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/CharP/LinearMaps.lean create mode 100644 Mathlib/FieldTheory/JacobsonNoether.lean diff --git a/Mathlib.lean b/Mathlib.lean index 55552ad2a0f9b..238444e50113b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -164,6 +164,7 @@ import Mathlib.Algebra.CharP.ExpChar import Mathlib.Algebra.CharP.IntermediateField import Mathlib.Algebra.CharP.Invertible import Mathlib.Algebra.CharP.Lemmas +import Mathlib.Algebra.CharP.LinearMaps import Mathlib.Algebra.CharP.LocalRing import Mathlib.Algebra.CharP.MixedCharZero import Mathlib.Algebra.CharP.Pi @@ -2946,6 +2947,7 @@ import Mathlib.FieldTheory.IsAlgClosed.Classification import Mathlib.FieldTheory.IsAlgClosed.Spectrum import Mathlib.FieldTheory.IsPerfectClosure import Mathlib.FieldTheory.IsSepClosed +import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension import Mathlib.FieldTheory.Laurent diff --git a/Mathlib/Algebra/Algebra/Basic.lean b/Mathlib/Algebra/Algebra/Basic.lean index 3e1c923709bb3..98aab66f635a1 100644 --- a/Mathlib/Algebra/Algebra/Basic.lean +++ b/Mathlib/Algebra/Algebra/Basic.lean @@ -133,18 +133,24 @@ end CommSemiring section Ring -variable [CommRing R] -variable (R) - /-- A `Semiring` that is an `Algebra` over a commutative ring carries a natural `Ring` structure. See note [reducible non-instances]. -/ -abbrev semiringToRing [Semiring A] [Algebra R A] : Ring A := +abbrev semiringToRing (R : Type*) [CommRing R] [Semiring A] [Algebra R A] : Ring A := { __ := (inferInstance : Semiring A) __ := Module.addCommMonoidToAddCommGroup R intCast := fun z => algebraMap R A z intCast_ofNat := fun z => by simp only [Int.cast_natCast, map_natCast] intCast_negSucc := fun z => by simp } +instance {R : Type*} [Ring R] : Algebra (Subring.center R) R where + toFun := Subtype.val + map_one' := rfl + map_mul' _ _ := rfl + map_zero' := rfl + map_add' _ _ := rfl + commutes' r x := (Subring.mem_center_iff.1 r.2 x).symm + smul_def' _ _ := rfl + end Ring end Algebra diff --git a/Mathlib/Algebra/CharP/LinearMaps.lean b/Mathlib/Algebra/CharP/LinearMaps.lean new file mode 100644 index 0000000000000..36f3baf18a80d --- /dev/null +++ b/Mathlib/Algebra/CharP/LinearMaps.lean @@ -0,0 +1,60 @@ +/- +Copyright (c) 2024 Wanyi He. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Wanyi He, Huanyu Zheng +-/ +import Mathlib.Algebra.CharP.Algebra +/-! +# Characteristic of the ring of linear Maps + +This file contains properties of the characteristic of the ring of linear maps. +The characteristic of the ring of linear maps is determined by its base ring. + +## Main Results + +- `CharP_if` : For a commutative semiring `R` and a `R`-module `M`, + the characteristic of `R` is equal to the characteristic of the `R`-linear + endomorphisms of `M` when `M` contains an element `x` such that + `r • x = 0` implies `r = 0`. + +## Notations + +- `R` is a commutative semiring +- `M` is a `R`-module + +## Implementation Notes + +One can also deduce similar result via `charP_of_injective_ringHom` and + `R → (M →ₗ[R] M) : r ↦ (fun (x : M) ↦ r • x)`. But this will require stronger condition + compared to `CharP_if`. + +-/ + +namespace Module + +variable {R M : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + +/-- For a commutative semiring `R` and a `R`-module `M`, if `M` contains an + element `x` such that `r • x = 0` implies `r = 0` (finding such element usually + depends on specific `•`), then the characteristic of `R` is equal to the + characteristic of the `R`-linear endomorphisms of `M`.-/ +theorem charP_end {p : ℕ} [hchar : CharP R p] + (hreduction : ∃ x : M, ∀ r : R, r • x = 0 → r = 0) : CharP (M →ₗ[R] M) p where + cast_eq_zero_iff' n := by + have exact : (n : M →ₗ[R] M) = (n : R) • 1 := by + simp only [Nat.cast_smul_eq_nsmul, nsmul_eq_mul, mul_one] + rw [exact, LinearMap.ext_iff, ← hchar.1] + exact ⟨fun h ↦ Exists.casesOn hreduction fun x hx ↦ hx n (h x), + fun h ↦ (congrArg (fun t ↦ ∀ x, t • x = 0) h).mpr fun x ↦ zero_smul R x⟩ + +end Module + +/-- For a division ring `D` with center `k`, the ring of `k`-linear endomorphisms + of `D` has the same characteristic as `D`-/ +instance {D : Type*} [DivisionRing D] {p : ℕ} [CharP D p] : + CharP (D →ₗ[(Subring.center D)] D) p := + charP_of_injective_ringHom (Algebra.lmul (Subring.center D) D).toRingHom.injective p + +instance {D : Type*} [DivisionRing D] {p : ℕ} [ExpChar D p] : + ExpChar (D →ₗ[Subring.center D] D) p := + expChar_of_injective_ringHom (Algebra.lmul (Subring.center D) D).toRingHom.injective p diff --git a/Mathlib/Algebra/CharP/Subring.lean b/Mathlib/Algebra/CharP/Subring.lean index e6d5faa37044a..23ed075efbeb1 100644 --- a/Mathlib/Algebra/CharP/Subring.lean +++ b/Mathlib/Algebra/CharP/Subring.lean @@ -3,8 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.Algebra.CharP.Defs -import Mathlib.Algebra.Ring.Subring.Defs +import Mathlib.Algebra.CharP.Algebra /-! # Characteristic of subrings @@ -33,4 +32,18 @@ instance subring (R : Type u) [Ring R] (p : ℕ) [CharP R p] (S : Subring R) : C instance subring' (R : Type u) [CommRing R] (p : ℕ) [CharP R p] (S : Subring R) : CharP S p := CharP.subring R p S +/-- The characteristic of a division ring is equal to the characteristic + of its center-/ +theorem charP_center_iff {R : Type u} [Ring R] {p : ℕ} : + CharP (Subring.center R) p ↔ CharP R p := + (algebraMap (Subring.center R) R).charP_iff Subtype.val_injective p + end CharP + +namespace ExpChar + +theorem expChar_center_iff {R : Type u} [Ring R] {p : ℕ} : + ExpChar (Subring.center R) p ↔ ExpChar R p := + (algebraMap (Subring.center R) R).expChar_iff Subtype.val_injective p + +end ExpChar diff --git a/Mathlib/Algebra/GroupWithZero/Conj.lean b/Mathlib/Algebra/GroupWithZero/Conj.lean index 17ccffc430192..6aea4c22972a7 100644 --- a/Mathlib/Algebra/GroupWithZero/Conj.lean +++ b/Mathlib/Algebra/GroupWithZero/Conj.lean @@ -14,9 +14,17 @@ assert_not_exists Multiset -- TODO -- assert_not_exists DenselyOrdered +namespace GroupWithZero + variable {α : Type*} [GroupWithZero α] {a b : α} @[simp] lemma isConj_iff₀ : IsConj a b ↔ ∃ c : α, c ≠ 0 ∧ c * a * c⁻¹ = b := by rw [IsConj, Units.exists_iff_ne_zero (p := (SemiconjBy · a b))] congr! 2 with c exact and_congr_right (mul_inv_eq_iff_eq_mul₀ · |>.symm) + +lemma conj_pow₀ {s : ℕ} {a d : α} (ha : a ≠ 0) : (a⁻¹ * d * a) ^ s = a⁻¹ * d ^ s * a := + let u : αˣ := ⟨a, a⁻¹, mul_inv_cancel₀ ha, inv_mul_cancel₀ ha⟩ + Units.conj_pow' u d s + +end GroupWithZero diff --git a/Mathlib/FieldTheory/JacobsonNoether.lean b/Mathlib/FieldTheory/JacobsonNoether.lean new file mode 100644 index 0000000000000..e019eaf9f4231 --- /dev/null +++ b/Mathlib/FieldTheory/JacobsonNoether.lean @@ -0,0 +1,195 @@ +/- +Copyright (c) 2024 F. Nuccio, H. Zheng, W. He, S. Wu, Y. Yuan, W. Jiao. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Filippo A. E. Nuccio, Huanyu Zheng, Sihan Wu, Wanyi He, Weichen Jiao, Yi Yuan +-/ +import Mathlib.Algebra.Central.Defs +import Mathlib.Algebra.CharP.LinearMaps +import Mathlib.Algebra.CharP.Subring +import Mathlib.Algebra.GroupWithZero.Conj +import Mathlib.Algebra.Lie.OfAssociative +import Mathlib.FieldTheory.PurelyInseparable + +/-! +# The Jacobson-Noether theorem + +This file contains a proof of the Jacobson-Noether theorem and some auxiliary lemmas. +Here we discuss different cases of characteristics of +the noncommutative division algebra `D` with center `k`. + +## Main Results + +- `exists_separable_and_not_isCentral` : (Jacobson-Noether theorem) For a + non-commutative algebraic division algebra `D` (with base ring + being its center `k`), then there exist an element `x` of + `D \ k` that is separable over its center. +- `exists_separable_and_not_isCentral'` : (Jacobson-Noether theorem) For a + non-commutative algebraic division algebra `D` (with base ring + being a field `L`), if the center of `D` over `L` is `L`, + then there exist an element `x` of `D \ L` that is separable over `L`. + +## Notations + +- `D` is a noncommutative division algebra +- `k` is the center of `D` + +## Implementation Notes + +Mathematically, `exists_separable_and_not_isCentral` and `exists_separable_and_not_isCentral'` +are equivalent. + +The difference however, is that the former takes `D` as the only variable +and fixing `D` would forces `k`. Whereas the later takes `D` and `L` as +separate variables constrained by certain relations. + +## Reference +* +-/ + +namespace JacobsonNoether + +variable {D : Type*} [DivisionRing D] [Algebra.IsAlgebraic (Subring.center D) D] + +local notation3 "k" => Subring.center D + +open Polynomial LinearMap LieAlgebra + +/-- If `D` is a purely inseparable extension of `k` with characteristic `p`, + then for every element `a` of `D`, there exists a natural number `n` + such that `a ^ (p ^ n)` is contained in `k`. -/ +lemma exists_pow_mem_center_of_inseparable (p : ℕ) [hchar : ExpChar D p] (a : D) + (hinsep : ∀ x : D, IsSeparable k x → x ∈ k) : ∃ n, a ^ (p ^ n) ∈ k := by + have := (@isPurelyInseparable_iff_pow_mem k D _ _ _ _ p (ExpChar.expChar_center_iff.2 hchar)).1 + have pure : IsPurelyInseparable k D := ⟨Algebra.IsAlgebraic.isIntegral, fun x hx ↦ by + rw [RingHom.mem_range, Subtype.exists] + exact ⟨x, ⟨hinsep x hx, rfl⟩⟩⟩ + obtain ⟨n, ⟨m, hm⟩⟩ := this pure a + have := Subalgebra.range_subset (R := k) ⟨(k).toSubsemiring, fun r ↦ r.2⟩ + exact ⟨n, Set.mem_of_subset_of_mem this <| Set.mem_range.2 ⟨m, hm⟩⟩ + +/-- If `D` is a purely inseparable extension of `k` with characteristic `p`, + then for every element `a` of `D \ k`, there exists a natural number `n` + **greater than 0** such that `a ^ (p ^ n)` is contained in `k`. -/ +lemma exists_pow_mem_center_of_inseparable' (p : ℕ) [ExpChar D p] {a : D} + (ha : a ∉ k) (hinsep : ∀ x : D, IsSeparable k x → x ∈ k) : ∃ n, 1 ≤ n ∧ a ^ (p ^ n) ∈ k := by + obtain ⟨n, hn⟩ := exists_pow_mem_center_of_inseparable p a hinsep + by_cases nzero : n = 0 + · rw [nzero, pow_zero, pow_one] at hn + exact (ha hn).elim + · exact ⟨n, ⟨Nat.one_le_iff_ne_zero.mpr nzero, hn⟩⟩ + +/-- If `D` is a purely inseparable extension of `k` of characteristic `p`, + then for every element `a` of `D \ k`, there exists a natural number `m` + greater than 0 such that `(a * x - x * a) ^ n = 0` (as linear maps) for + every `n` greater than `(p ^ m)`. -/ +lemma exist_pow_eq_zero_of_le (p : ℕ) [hchar : ExpChar D p] + {a : D} (ha : a ∉ k) (hinsep : ∀ x : D, IsSeparable k x → x ∈ k): + ∃ m, 1 ≤ m ∧ ∀ n, p ^ m ≤ n → (ad k D a)^[n] = 0 := by + obtain ⟨m, hm⟩ := exists_pow_mem_center_of_inseparable' p ha hinsep + refine ⟨m, ⟨hm.1, fun n hn ↦ ?_⟩⟩ + have inter : (ad k D a)^[p ^ m] = 0 := by + ext x + rw [ad_eq_lmul_left_sub_lmul_right, ← pow_apply, Pi.sub_apply, + sub_pow_expChar_pow_of_commute p m (commute_mulLeft_right a a), sub_apply, + pow_mulLeft, mulLeft_apply, pow_mulRight, mulRight_apply, Pi.zero_apply, + Subring.mem_center_iff.1 hm.2 x] + exact sub_eq_zero_of_eq rfl + rw [(Nat.sub_eq_iff_eq_add hn).1 rfl, Function.iterate_add, inter, Pi.comp_zero, + iterate_map_zero, Function.const_zero] + +variable (D) in +/-- Jacobson-Noether theorem: For a non-commutative division algebra + `D` that is algebraic over its center `k`, there exists an element + `x` of `D \ k` that is separable over `k`. -/ +theorem exists_separable_and_not_isCentral (H : k ≠ (⊤ : Subring D)) : + ∃ x : D, x ∉ k ∧ IsSeparable k x := by + obtain ⟨p, hp⟩ := ExpChar.exists D + by_contra! insep + replace insep : ∀ x : D, IsSeparable k x → x ∈ k := + fun x h ↦ Classical.byContradiction fun hx ↦ insep x hx h + -- The element `a` below is in `D` but not in `k`. + obtain ⟨a, ha⟩ := not_forall.mp <| mt (Subring.eq_top_iff' k).mpr H + have ha₀ : a ≠ 0 := fun nh ↦ nh ▸ ha <| Subring.zero_mem k + -- We construct another element `b` that does not commute with `a`. + obtain ⟨b, hb1⟩ : ∃ b : D , ad k D a b ≠ 0 := by + rw [Subring.mem_center_iff, not_forall] at ha + use ha.choose + show a * ha.choose - ha.choose * a ≠ 0 + simpa only [ne_eq, sub_eq_zero] using Ne.symm ha.choose_spec + -- We find a maximum natural number `n` such that `(a * x - x * a) ^ n b ≠ 0`. + obtain ⟨n, hn, hb⟩ : ∃ n, 0 < n ∧ (ad k D a)^[n] b ≠ 0 ∧ (ad k D a)^[n + 1] b = 0 := by + obtain ⟨m, -, hm2⟩ := exist_pow_eq_zero_of_le p ha insep + have h_exist : ∃ n, 0 < n ∧ (ad k D a)^[n + 1] b = 0 := ⟨p ^ m, + ⟨expChar_pow_pos D p m, by rw [hm2 (p ^ m + 1) (Nat.le_add_right _ _)]; rfl⟩⟩ + classical + refine ⟨Nat.find h_exist, ⟨(Nat.find_spec h_exist).1, ?_, (Nat.find_spec h_exist).2⟩⟩ + set t := (Nat.find h_exist - 1 : ℕ) with ht + by_cases h_pos : 0 < t + · convert (ne_eq _ _) ▸ not_and.mp (Nat.find_min h_exist (m := t) (by omega)) h_pos + omega + · suffices h_find: Nat.find h_exist = 1 by + rwa [h_find] + rw [not_lt, Nat.le_zero, ht, Nat.sub_eq_zero_iff_le] at h_pos + linarith [(Nat.find_spec h_exist).1] + -- We define `c` to be the value that we proved above to be non-zero. + set c := (ad k D a)^[n] b with hc_def + let _ : Invertible c := ⟨c⁻¹, inv_mul_cancel₀ hb.1, mul_inv_cancel₀ hb.1⟩ + -- We prove that `c` commutes with `a`. + have hc : a * c = c * a := by + apply eq_of_sub_eq_zero + rw [← mulLeft_apply (R := k), ← mulRight_apply (R := k)] + suffices ad k D a c = 0 from by + rw [← this]; rfl + rw [← Function.iterate_succ_apply' (ad k D a) n b, hb.2] + -- We now make some computation to obtain the final equation. + set d := c⁻¹ * a * (ad k D a)^[n - 1] b with hd_def + have hc': c⁻¹ * a = a * c⁻¹ := by + apply_fun (c⁻¹ * · * c⁻¹) at hc + rw [mul_assoc, mul_assoc, mul_inv_cancel₀ hb.1, mul_one, ← mul_assoc, + inv_mul_cancel₀ hb.1, one_mul] at hc + exact hc + have c_eq : a * (ad k D a)^[n - 1] b - (ad k D a)^[n - 1] b * a = c := by + rw [hc_def, ← Nat.sub_add_cancel hn, Function.iterate_succ_apply' (ad k D a) _ b]; rfl + have eq1 : c⁻¹ * a * (ad k D a)^[n - 1] b - c⁻¹ * (ad k D a)^[n - 1] b * a = 1 := by + simp_rw [mul_assoc, (mul_sub_left_distrib c⁻¹ _ _).symm, c_eq, inv_mul_cancel_of_invertible] + -- We show that `a` commutes with `d`. + have deq : a * d - d * a = a := by + nth_rw 3 [← mul_one a] + rw [hd_def, ← eq1, mul_sub, mul_assoc _ _ a, sub_right_inj, hc', + ← mul_assoc, ← mul_assoc, ← mul_assoc] + -- This then yields a contradiction. + apply_fun (a⁻¹ * · ) at deq + rw [mul_sub, ← mul_assoc, inv_mul_cancel₀ ha₀, one_mul, ← mul_assoc, sub_eq_iff_eq_add] at deq + obtain ⟨r, hr⟩ := exists_pow_mem_center_of_inseparable p d insep + apply_fun (· ^ (p ^ r)) at deq + rw [add_pow_expChar_pow_of_commute p r (Commute.one_left _) , one_pow, + GroupWithZero.conj_pow₀ ha₀, ← hr.comm, mul_assoc, inv_mul_cancel₀ ha₀, mul_one, + self_eq_add_left] at deq + exact one_ne_zero deq + +open Subring Algebra in +/-- Jacobson-Noether theorem: For a non-commutative division algebra `D` + that is algebraic over a field `L`, if the center of + `D` coincides with `L`, then there exist an element `x` of `D \ L` + that is separable over `L`. -/ +theorem exists_separable_and_not_isCentral' {L D : Type*} [Field L] [DivisionRing D] + [Algebra L D] [Algebra.IsAlgebraic L D] [Algebra.IsCentral L D] + (hneq : (⊥ : Subalgebra L D) ≠ ⊤) : + ∃ x : D, x ∉ (⊥ : Subalgebra L D) ∧ IsSeparable L x := by + have hcenter : Subalgebra.center L D = ⊥ := le_bot_iff.mp IsCentral.out + have ntrivial : Subring.center D ≠ ⊤ := + congr(Subalgebra.toSubring $hcenter).trans_ne (Subalgebra.toSubring_injective.ne hneq) + set φ := Subalgebra.equivOfEq (⊥ : Subalgebra L D) (.center L D) hcenter.symm + set equiv : L ≃+* (center D) := ((botEquiv L D).symm.trans φ).toRingEquiv + let _ : Algebra L (center D) := equiv.toRingHom.toAlgebra + let _ : Algebra (center D) L := equiv.symm.toRingHom.toAlgebra + have _ : IsScalarTower L (center D) D := .of_algebraMap_eq fun _ ↦ rfl + have _ : IsScalarTower (center D) L D := .of_algebraMap_eq fun x ↦ by + rw [IsScalarTower.algebraMap_apply L (center D)] + congr + exact (equiv.apply_symm_apply x).symm + have _ : Algebra.IsAlgebraic (center D) D := .tower_top (K := L) _ + obtain ⟨x, hxd, hx⟩ := exists_separable_and_not_isCentral D ntrivial + exact ⟨x, ⟨by rwa [← Subalgebra.center_toSubring L, hcenter] at hxd, IsSeparable.tower_top _ hx⟩⟩ + +end JacobsonNoether From be959a9ade921f874584eb5761e7c4f6c82f6651 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 22 Nov 2024 13:33:56 +0000 Subject: [PATCH 036/250] feat(LinearAlgebra/Matrix/Determinant/TotallyUnimodular): iff_fintype (#19366) add a lemma relating the definition of `Matrix.IsTotallyUnimodular` to `Fintype` rather than `Fin`. Co-authored-by: blizzard_inc --- .../Matrix/Determinant/TotallyUnimodular.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean index cd7b6cae09bfd..03da13b5d1757 100644 --- a/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean +++ b/Mathlib/LinearAlgebra/Matrix/Determinant/TotallyUnimodular.lean @@ -54,6 +54,18 @@ lemma isTotallyUnimodular_iff (A : Matrix m n R) : A.IsTotallyUnimodular ↔ · intro _ _ _ _ _ apply hA +lemma isTotallyUnimodular_iff_fintype.{w} (A : Matrix m n R) : A.IsTotallyUnimodular ↔ + ∀ (ι : Type w) [Fintype ι] [DecidableEq ι], ∀ f : ι → m, ∀ g : ι → n, + (A.submatrix f g).det ∈ Set.range SignType.cast := by + rw [isTotallyUnimodular_iff] + constructor + · intro hA ι _ _ f g + specialize hA (Fintype.card ι) (f ∘ (Fintype.equivFin ι).symm) (g ∘ (Fintype.equivFin ι).symm) + rwa [←submatrix_submatrix, det_submatrix_equiv_self] at hA + · intro hA k f g + specialize hA (ULift (Fin k)) (f ∘ Equiv.ulift) (g ∘ Equiv.ulift) + rwa [←submatrix_submatrix, det_submatrix_equiv_self] at hA + lemma IsTotallyUnimodular.apply {A : Matrix m n R} (hA : A.IsTotallyUnimodular) (i : m) (j : n) : A i j ∈ Set.range SignType.cast := by From e03117db88529135436883761084450509f14d15 Mon Sep 17 00:00:00 2001 From: damiano Date: Fri, 22 Nov 2024 17:25:12 +0000 Subject: [PATCH 037/250] perf: header linter performs fewer checks (#19260) Trying to recover from the [slowdown](http://speed.lean-fro.org/mathlib4/run-detail/e7b27246-a3e6-496a-b552-ff4b45c7236e/d15f04e158ee396b59b6d37cea21d2ebba0dcf7d) introduced in #18033. --- Mathlib/Tactic/Linter/Header.lean | 21 +++++++++++++++++++-- 1 file changed, 19 insertions(+), 2 deletions(-) diff --git a/Mathlib/Tactic/Linter/Header.lean b/Mathlib/Tactic/Linter/Header.lean index 9f5be314fe2ff..19b353da35598 100644 --- a/Mathlib/Tactic/Linter/Header.lean +++ b/Mathlib/Tactic/Linter/Header.lean @@ -218,6 +218,15 @@ def isInMathlib (modName : Name) : IO Bool := do return (ml.map (·.module == modName)).any (·) else return false +/-- `inMathlibRef` is +* `none` at initialization time; +* `some true` if the `header` linter has already discovered that the current file + is imported in `Mathlib.lean`; +* `some false` if the `header` linter has already discovered that the current file + is *not* imported in `Mathlib.lean`. +-/ +initialize inMathlibRef : IO.Ref (Option Bool) ← IO.mkRef none + /-- The "header" style linter checks that a file starts with ``` @@ -290,9 +299,17 @@ def duplicateImportsCheck (imports : Array Syntax) : CommandElabM Unit := do @[inherit_doc Mathlib.Linter.linter.style.header] def headerLinter : Linter where run := withSetOptionIn fun stx ↦ do let mainModule ← getMainModule + let inMathlib? := ← match ← inMathlibRef.get with + | some d => return d + | none => do + let val ← isInMathlib mainModule + -- We store the answer to the question "is this file in `Mathlib.lean`?" in `inMathlibRef` + -- to avoid recomputing its value on every command. This is a performance optimisation. + inMathlibRef.set (some val) + return val -- The linter skips files not imported in `Mathlib.lean`, to avoid linting "scratch files". - -- However, it is active in the test file `MathlibTest.Header` for the linter itself. - unless (← isInMathlib mainModule) || mainModule == `MathlibTest.Header do return + -- It is however active in the test file `MathlibTest.Header` for the linter itself. + unless inMathlib? || mainModule == `MathlibTest.Header do return unless Linter.getLinterValue linter.style.header (← getOptions) do return if (← get).messages.hasErrors then From 178af37e20ade7b888c1ceb1edbc00dcc331562b Mon Sep 17 00:00:00 2001 From: FMLJohn <417963616@qq.com> Date: Fri, 22 Nov 2024 18:05:10 +0000 Subject: [PATCH 038/250] feat(RingTheory/PowerSeries/Basic): `Polynomial.coe_sub` and `Polynomial.coe_neg` (#19339) These are analogous to the existing `Polynomial.coe_add` and `Polynomial.coe_zero`, about the coercion from `Polynomial` to `PowerSeries`. [Zulip thread](https://leanprover.zulipchat.com/#narrow/channel/217875-Is-there-code-for-X.3F/topic/.60.281.20-.20X.20.3A.20.E2.84.A4.5BX.5D.29.2EtoPowerSeries.20.3D.201.20-.20.28PowerSeries.2EX.20.3A.20.E2.84.A4.E2.9F.A6X.E2.9F.A7.29.60/near/483708229). --- Mathlib/RingTheory/LaurentSeries.lean | 10 +++++----- Mathlib/RingTheory/PowerSeries/Basic.lean | 16 ++++++++++++++++ 2 files changed, 21 insertions(+), 5 deletions(-) diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index d543b3fa95654..f8ee85f00a0c0 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -933,7 +933,7 @@ theorem exists_ratFunc_val_lt (f : K⸨X⸩) (γ : ℤₘ₀ˣ) : · erw [hs, ← F_mul, PowerSeries.coe_pow, PowerSeries.coe_X, RatFunc.coe_mul, zpow_neg, zpow_ofNat, inv_eq_one_div (RatFunc.X ^ s), RatFunc.coe_div, RatFunc.coe_pow, RatFunc.coe_X, RatFunc.coe_one, ← inv_eq_one_div, ← mul_sub, map_mul, map_inv₀, ← PowerSeries.coe_X, - valuation_X_pow, ← hs, ← RatFunc.coe_coe, ← coe_sub, ← coe_algebraMap, + valuation_X_pow, ← hs, ← RatFunc.coe_coe, ← PowerSeries.coe_sub, ← coe_algebraMap, valuation_of_algebraMap, ← Units.val_mk0 (a := ((Multiplicative.ofAdd f.order : Multiplicative ℤ) : ℤₘ₀)), ← hη] apply inv_mul_lt_of_lt_mul₀ @@ -943,8 +943,8 @@ theorem exists_ratFunc_val_lt (f : K⸨X⸩) (γ : ℤₘ₀ˣ) : · obtain ⟨s, hs⟩ := Int.exists_eq_neg_ofNat (Int.neg_nonpos_of_nonneg (not_lt.mp ord_nonpos)) obtain ⟨P, hP⟩ := exists_Polynomial_intValuation_lt (PowerSeries.X ^ s * F) γ use P - erw [← X_order_mul_powerSeriesPart (neg_inj.1 hs).symm, ← RatFunc.coe_coe, ← coe_sub, - ← coe_algebraMap, valuation_of_algebraMap] + erw [← X_order_mul_powerSeriesPart (neg_inj.1 hs).symm, ← RatFunc.coe_coe, + ← PowerSeries.coe_sub, ← coe_algebraMap, valuation_of_algebraMap] exact hP theorem coe_range_dense : DenseRange ((↑) : RatFunc K → K⸨X⸩) := by @@ -981,7 +981,7 @@ theorem inducing_coe : IsUniformInducing ((↑) : RatFunc K → K⸨X⸩) := by refine ⟨⟨d, by rfl⟩, subset_trans (fun _ _ ↦ pre_R ?_) pre_T⟩ apply hd simp only [sub_zero, Set.mem_setOf_eq] - erw [← coe_sub, ← valuation_eq_LaurentSeries_valuation] + erw [← RatFunc.coe_sub, ← valuation_eq_LaurentSeries_valuation] assumption · rintro ⟨_, ⟨hT, pre_T⟩⟩ obtain ⟨d, hd⟩ := Valued.mem_nhds.mp hT @@ -993,7 +993,7 @@ theorem inducing_coe : IsUniformInducing ((↑) : RatFunc K → K⸨X⸩) := by · refine subset_trans (fun _ _ ↦ ?_) pre_T apply hd erw [Set.mem_setOf_eq, sub_zero, valuation_eq_LaurentSeries_valuation, - coe_sub] + RatFunc.coe_sub] assumption theorem continuous_coe : Continuous ((↑) : RatFunc K → K⸨X⸩) := diff --git a/Mathlib/RingTheory/PowerSeries/Basic.lean b/Mathlib/RingTheory/PowerSeries/Basic.lean index 6030481adc2e7..85884f5a24154 100644 --- a/Mathlib/RingTheory/PowerSeries/Basic.lean +++ b/Mathlib/RingTheory/PowerSeries/Basic.lean @@ -758,6 +758,7 @@ namespace Polynomial open Finsupp Polynomial +section CommSemiring variable {R : Type*} [CommSemiring R] (φ ψ : R[X]) -- Porting note: added so we can add the `@[coe]` attribute @@ -877,6 +878,21 @@ theorem coeToPowerSeries.algHom_apply : coeToPowerSeries.algHom A φ = PowerSeries.map (algebraMap R A) ↑φ := rfl +end CommSemiring + +section CommRing +variable {R : Type*} [CommRing R] + +@[simp, norm_cast] +lemma coe_neg (p : R[X]) : ((- p : R[X]) : PowerSeries R) = - p := + coeToPowerSeries.ringHom.map_neg p + +@[simp, norm_cast] +lemma coe_sub (p q : R[X]) : ((p - q : R[X]) : PowerSeries R) = p - q := + coeToPowerSeries.ringHom.map_sub p q + +end CommRing + end Polynomial namespace PowerSeries From 3d401e0bb3b0295605d8083d5683efd6dbbea2eb Mon Sep 17 00:00:00 2001 From: Gio <102917377+gio256@users.noreply.github.com> Date: Fri, 22 Nov 2024 18:29:55 +0000 Subject: [PATCH 039/250] feat(AlgebraicTopology/SimplicialSet): Add auxiliary ext lemma for paths (#19349) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit We implement the generalization of the path extensionality lemmas suggested by @joelriou in #19057. Two paths of the same nonzero length can be identified by constructing an identification between each of their arrows. Co-Authored-By: [Joël Riou](https://github.com/joelriou) --- Mathlib/AlgebraicTopology/SimplicialSet/Path.lean | 12 ++++++++++++ 1 file changed, 12 insertions(+) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 696d2c7aa9498..047e74c18ff93 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -73,4 +73,16 @@ lemma spine_map_subinterval {n : ℕ} (j l : ℕ) (hjl : j + l ≤ n) (Δ : X _[ · simp only [spine_arrow, Path.interval, ← FunctorToTypes.map_comp_apply, ← op_comp, mkOfSucc_subinterval_eq] +/-- Two paths of the same nonzero length are equal if all of their arrows are equal. -/ +@[ext] +lemma Path.ext' {n : ℕ} {f g : Path X (n + 1)} + (h : ∀ i : Fin (n + 1), f.arrow i = g.arrow i) : + f = g := by + ext j + · rcases Fin.eq_castSucc_or_eq_last j with ⟨k, hk⟩ | hl + · rw [hk, ← f.arrow_src k, ← g.arrow_src k, h] + · simp only [hl, ← Fin.succ_last] + rw [← f.arrow_tgt (Fin.last n), ← g.arrow_tgt (Fin.last n), h] + · exact h j + end SSet From dd74f0185f0bdda192392376744bfd1a7a9d977b Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Fri, 22 Nov 2024 21:24:50 +0000 Subject: [PATCH 040/250] chore: split Algebra.Group.Nat (#19375) --- Mathlib.lean | 5 +- Mathlib/Algebra/Group/Int.lean | 3 +- Mathlib/Algebra/Group/Nat/Basic.lean | 65 +++++++++++++++ .../Algebra/Group/{Nat.lean => Nat/Even.lean} | 81 +------------------ Mathlib/Algebra/Group/Nat/TypeTags.lean | 24 ++++++ Mathlib/Algebra/Group/Nat/Units.lean | 38 +++++++++ .../Algebra/Group/Submonoid/Operations.lean | 3 +- Mathlib/Algebra/GroupPower/IterateHom.lean | 1 - Mathlib/Algebra/Order/Group/Nat.lean | 2 +- Mathlib/Algebra/Order/Ring/Basic.lean | 1 + Mathlib/Algebra/Ring/Nat.lean | 4 +- Mathlib/Algebra/Ring/Parity.lean | 1 + Mathlib/Data/List/SplitLengths.lean | 2 +- Mathlib/Data/Multiset/Basic.lean | 7 +- Mathlib/Data/Nat/Bits.lean | 3 +- Mathlib/Data/Nat/Bitwise.lean | 2 +- Mathlib/Data/Nat/Cast/Basic.lean | 3 +- Mathlib/Data/Nat/GCD/Basic.lean | 1 + Mathlib/Data/Nat/PSub.lean | 2 +- Mathlib/Data/Nat/Prime/Defs.lean | 2 + Mathlib/Data/Nat/Size.lean | 1 + Mathlib/Data/Set/Enumerate.lean | 2 +- .../SpecificGroups/Quaternion.lean | 1 - Mathlib/Tactic/Sat/FromLRAT.lean | 3 +- MathlibTest/levenshtein.lean | 2 +- 25 files changed, 159 insertions(+), 100 deletions(-) create mode 100644 Mathlib/Algebra/Group/Nat/Basic.lean rename Mathlib/Algebra/Group/{Nat.lean => Nat/Even.lean} (60%) create mode 100644 Mathlib/Algebra/Group/Nat/TypeTags.lean create mode 100644 Mathlib/Algebra/Group/Nat/Units.lean diff --git a/Mathlib.lean b/Mathlib.lean index 238444e50113b..fc3aa1d9991a6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -281,7 +281,10 @@ import Mathlib.Algebra.Group.Int import Mathlib.Algebra.Group.Invertible.Basic import Mathlib.Algebra.Group.Invertible.Defs import Mathlib.Algebra.Group.MinimalAxioms -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.Nat.Even +import Mathlib.Algebra.Group.Nat.TypeTags +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.NatPowAssoc import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.PNatPowAssoc diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 1fcf3f10ef98c..948531494a03b 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -3,7 +3,8 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Even +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.Units.Basic import Mathlib.Data.Int.Sqrt diff --git a/Mathlib/Algebra/Group/Nat/Basic.lean b/Mathlib/Algebra/Group/Nat/Basic.lean new file mode 100644 index 0000000000000..75212737a122b --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/Basic.lean @@ -0,0 +1,65 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Defs + +/-! +# The natural numbers form a monoid + +This file contains the additive and multiplicative monoid instances on the natural numbers. + +See note [foundational algebra order theory]. +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +namespace Nat + +/-! ### Instances -/ + +instance instAddCancelCommMonoid : AddCancelCommMonoid ℕ where + add := Nat.add + add_assoc := Nat.add_assoc + zero := Nat.zero + zero_add := Nat.zero_add + add_zero := Nat.add_zero + add_comm := Nat.add_comm + nsmul m n := m * n + nsmul_zero := Nat.zero_mul + nsmul_succ := succ_mul + add_left_cancel _ _ _ := Nat.add_left_cancel + +instance instCommMonoid : CommMonoid ℕ where + mul := Nat.mul + mul_assoc := Nat.mul_assoc + one := Nat.succ Nat.zero + one_mul := Nat.one_mul + mul_one := Nat.mul_one + mul_comm := Nat.mul_comm + npow m n := n ^ m + npow_zero := Nat.pow_zero + npow_succ _ _ := rfl + +/-! +### Extra instances to short-circuit type class resolution + +These also prevent non-computable instances being used to construct these instances non-computably. +-/ + +instance instAddCommMonoid : AddCommMonoid ℕ := by infer_instance +instance instAddMonoid : AddMonoid ℕ := by infer_instance +instance instMonoid : Monoid ℕ := by infer_instance +instance instCommSemigroup : CommSemigroup ℕ := by infer_instance +instance instSemigroup : Semigroup ℕ := by infer_instance +instance instAddCommSemigroup : AddCommSemigroup ℕ := by infer_instance +instance instAddSemigroup : AddSemigroup ℕ := by infer_instance + +/-! ### Miscellaneous lemmas -/ + +-- We set the simp priority slightly lower than default; later more general lemmas will replace it. +@[simp 900] protected lemma nsmul_eq_mul (m n : ℕ) : m • n = m * n := rfl + +end Nat diff --git a/Mathlib/Algebra/Group/Nat.lean b/Mathlib/Algebra/Group/Nat/Even.lean similarity index 60% rename from Mathlib/Algebra/Group/Nat.lean rename to Mathlib/Algebra/Group/Nat/Even.lean index 42bd89df34d7b..86fcfd3209c0d 100644 --- a/Mathlib/Algebra/Group/Nat.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -4,76 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.Group.Even -import Mathlib.Algebra.Group.Units.Defs +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Data.Nat.Sqrt /-! -# The natural numbers form a monoid - -This file contains the additive and multiplicative monoid instances on the natural numbers. - -See note [foundational algebra order theory]. +# `IsSquare` and `Even` for natural numbers -/ assert_not_exists MonoidWithZero assert_not_exists DenselyOrdered -open Multiplicative - namespace Nat -/-! ### Instances -/ - -instance instAddCancelCommMonoid : AddCancelCommMonoid ℕ where - add := Nat.add - add_assoc := Nat.add_assoc - zero := Nat.zero - zero_add := Nat.zero_add - add_zero := Nat.add_zero - add_comm := Nat.add_comm - nsmul m n := m * n - nsmul_zero := Nat.zero_mul - nsmul_succ := succ_mul - add_left_cancel _ _ _ := Nat.add_left_cancel - -instance instCommMonoid : CommMonoid ℕ where - mul := Nat.mul - mul_assoc := Nat.mul_assoc - one := Nat.succ Nat.zero - one_mul := Nat.one_mul - mul_one := Nat.mul_one - mul_comm := Nat.mul_comm - npow m n := n ^ m - npow_zero := Nat.pow_zero - npow_succ _ _ := rfl - -/-! -### Extra instances to short-circuit type class resolution - -These also prevent non-computable instances being used to construct these instances non-computably. --/ - -instance instAddCommMonoid : AddCommMonoid ℕ := by infer_instance -instance instAddMonoid : AddMonoid ℕ := by infer_instance -instance instMonoid : Monoid ℕ := by infer_instance -instance instCommSemigroup : CommSemigroup ℕ := by infer_instance -instance instSemigroup : Semigroup ℕ := by infer_instance -instance instAddCommSemigroup : AddCommSemigroup ℕ := by infer_instance -instance instAddSemigroup : AddSemigroup ℕ := by infer_instance - -/-! ### Miscellaneous lemmas -/ - --- We set the simp priority slightly lower than default; later more general lemmas will replace it. -@[simp 900] protected lemma nsmul_eq_mul (m n : ℕ) : m • n = m * n := rfl - -section Multiplicative - -lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ - -@[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm - -end Multiplicative - /-! #### Parity -/ variable {m n : ℕ} @@ -158,23 +100,4 @@ example (m n : ℕ) (h : Even m) : ¬Even (n + 3) ↔ Even (m ^ 2 + m + n) := by -- Porting note: the `simp` lemmas about `bit*` no longer apply. example : ¬Even 25394535 := by decide -/-! #### Units -/ - -lemma units_eq_one (u : ℕˣ) : u = 1 := Units.ext <| Nat.eq_one_of_dvd_one ⟨u.inv, u.val_inv.symm⟩ - -lemma addUnits_eq_zero (u : AddUnits ℕ) : u = 0 := - AddUnits.ext <| (Nat.eq_zero_of_add_eq_zero u.val_neg).1 - -@[simp] protected lemma isUnit_iff {n : ℕ} : IsUnit n ↔ n = 1 where - mp := by rintro ⟨u, rfl⟩; obtain rfl := Nat.units_eq_one u; rfl - mpr h := h.symm ▸ ⟨1, rfl⟩ - -instance unique_units : Unique ℕˣ where - default := 1 - uniq := Nat.units_eq_one - -instance unique_addUnits : Unique (AddUnits ℕ) where - default := 0 - uniq := Nat.addUnits_eq_zero - end Nat diff --git a/Mathlib/Algebra/Group/Nat/TypeTags.lean b/Mathlib/Algebra/Group/Nat/TypeTags.lean new file mode 100644 index 0000000000000..0efe1e8dc736b --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/TypeTags.lean @@ -0,0 +1,24 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.TypeTags.Basic + +/-! +# Lemmas about `Multiplicative ℕ` +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +open Multiplicative + +namespace Nat + +lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ + +@[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm + +end Nat diff --git a/Mathlib/Algebra/Group/Nat/Units.lean b/Mathlib/Algebra/Group/Nat/Units.lean new file mode 100644 index 0000000000000..edc5970d8d2e4 --- /dev/null +++ b/Mathlib/Algebra/Group/Nat/Units.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro +-/ +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Algebra.Group.Units.Defs +import Mathlib.Logic.Unique + +/-! +# The unit of the natural numbers +-/ + +assert_not_exists MonoidWithZero +assert_not_exists DenselyOrdered + +namespace Nat + +/-! #### Units -/ + +lemma units_eq_one (u : ℕˣ) : u = 1 := Units.ext <| Nat.eq_one_of_dvd_one ⟨u.inv, u.val_inv.symm⟩ + +lemma addUnits_eq_zero (u : AddUnits ℕ) : u = 0 := + AddUnits.ext <| (Nat.eq_zero_of_add_eq_zero u.val_neg).1 + +@[simp] protected lemma isUnit_iff {n : ℕ} : IsUnit n ↔ n = 1 where + mp := by rintro ⟨u, rfl⟩; obtain rfl := Nat.units_eq_one u; rfl + mpr h := h.symm ▸ ⟨1, rfl⟩ + +instance unique_units : Unique ℕˣ where + default := 1 + uniq := Nat.units_eq_one + +instance unique_addUnits : Unique (AddUnits ℕ) where + default := 0 + uniq := Nat.addUnits_eq_zero + +end Nat diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 46c96865817f9..36ee592c01e9d 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -5,9 +5,10 @@ Authors: Johannes Hölzl, Kenny Lau, Johan Commelin, Mario Carneiro, Kevin Buzza Amelia Livingston, Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Faithful -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Submonoid.Basic +import Mathlib.Algebra.Group.TypeTags.Basic /-! # Operations on `Submonoid`s diff --git a/Mathlib/Algebra/GroupPower/IterateHom.lean b/Mathlib/Algebra/GroupPower/IterateHom.lean index 1a3170e11ac09..ae097ed297d67 100644 --- a/Mathlib/Algebra/GroupPower/IterateHom.lean +++ b/Mathlib/Algebra/GroupPower/IterateHom.lean @@ -5,7 +5,6 @@ Authors: Yury Kudryashov -/ import Mathlib.Algebra.Group.Action.Opposite import Mathlib.Algebra.Group.Int -import Mathlib.Algebra.Group.Nat import Mathlib.Logic.Function.Iterate import Mathlib.Tactic.Common diff --git a/Mathlib/Algebra/Order/Group/Nat.lean b/Mathlib/Algebra/Order/Group/Nat.lean index 5d7b7628327fd..1385365e53779 100644 --- a/Mathlib/Algebra/Order/Group/Nat.lean +++ b/Mathlib/Algebra/Order/Group/Nat.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Order.Monoid.Canonical.Defs import Mathlib.Algebra.Order.Sub.Defs import Mathlib.Data.Nat.Defs diff --git a/Mathlib/Algebra/Order/Ring/Basic.lean b/Mathlib/Algebra/Order/Ring/Basic.lean index 65844e24d7f30..6f6a50a96337c 100644 --- a/Mathlib/Algebra/Order/Ring/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2015 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis -/ +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Algebra.Order.Ring.Defs import Mathlib.Algebra.Ring.Parity diff --git a/Mathlib/Algebra/Ring/Nat.lean b/Mathlib/Algebra/Ring/Nat.lean index 9b456d31dec8a..e2c57fc380685 100644 --- a/Mathlib/Algebra/Ring/Nat.lean +++ b/Mathlib/Algebra/Ring/Nat.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Mathlib.Algebra.CharZero.Defs -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Ring.Defs /-! @@ -15,8 +15,6 @@ This file contains the commutative semiring instance on the natural numbers. See note [foundational algebra order theory]. -/ -open Multiplicative - namespace Nat /-! ### Instances -/ diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 018e1506eb3ec..ff8df726aad0a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -3,6 +3,7 @@ Copyright (c) 2022 Damiano Testa. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa -/ +import Mathlib.Algebra.Group.Nat.Even import Mathlib.Data.Nat.Cast.Basic import Mathlib.Data.Nat.Cast.Commute import Mathlib.Data.Set.Operations diff --git a/Mathlib/Data/List/SplitLengths.lean b/Mathlib/Data/List/SplitLengths.lean index e41894fe16e0c..ae9a0052291fa 100644 --- a/Mathlib/Data/List/SplitLengths.lean +++ b/Mathlib/Data/List/SplitLengths.lean @@ -3,8 +3,8 @@ Copyright (c) 2024 Daniel Weber. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Daniel Weber -/ +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Order.MinMax -import Mathlib.Algebra.Group.Nat /-! # Splitting a list to chunks of specified lengths diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index a594866b4536d..49f764b115ba5 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -3,13 +3,14 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Hom.Defs +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Algebra.Order.Sub.Unbundled.Basic +import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Perm.Lattice import Mathlib.Data.List.Perm.Subperm import Mathlib.Data.Set.List import Mathlib.Order.Hom.Basic -import Mathlib.Data.List.Perm.Lattice -import Mathlib.Data.List.Perm.Basic /-! # Multisets diff --git a/Mathlib/Data/Nat/Bits.lean b/Mathlib/Data/Nat/Bits.lean index 7079602da88d9..fcc0bc2b46077 100644 --- a/Mathlib/Data/Nat/Bits.lean +++ b/Mathlib/Data/Nat/Bits.lean @@ -3,8 +3,7 @@ Copyright (c) 2022 Praneeth Kolichala. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Praneeth Kolichala -/ -import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Data.Nat.Defs import Mathlib.Data.Nat.BinaryRec import Mathlib.Data.List.Defs diff --git a/Mathlib/Data/Nat/Bitwise.lean b/Mathlib/Data/Nat/Bitwise.lean index 2de92522a8e9a..e3b32b271b950 100644 --- a/Mathlib/Data/Nat/Bitwise.lean +++ b/Mathlib/Data/Nat/Bitwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel, Alex Keizer -/ -import Mathlib.Algebra.Group.Units.Basic +import Mathlib.Algebra.Group.Nat.Even import Mathlib.Algebra.NeZero import Mathlib.Algebra.Ring.Nat import Mathlib.Data.List.GetD diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index b8cda137de314..d9fd72598d7bf 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Divisibility.Basic +import Mathlib.Algebra.Group.Even +import Mathlib.Algebra.Group.TypeTags.Hom import Mathlib.Algebra.Ring.Hom.Defs import Mathlib.Algebra.Ring.Nat -import Mathlib.Algebra.Group.TypeTags.Hom /-! # Cast of natural numbers (additional theorems) diff --git a/Mathlib/Data/Nat/GCD/Basic.lean b/Mathlib/Data/Nat/GCD/Basic.lean index ff926c64b4005..2adab886eac34 100644 --- a/Mathlib/Data/Nat/GCD/Basic.lean +++ b/Mathlib/Data/Nat/GCD/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura -/ import Batteries.Data.Nat.Gcd +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.Ring.Nat diff --git a/Mathlib/Data/Nat/PSub.lean b/Mathlib/Data/Nat/PSub.lean index e4d0b1efeaaae..dce691789a710 100644 --- a/Mathlib/Data/Nat/PSub.lean +++ b/Mathlib/Data/Nat/PSub.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic /-! # Partial predecessor and partial subtraction on the natural numbers diff --git a/Mathlib/Data/Nat/Prime/Defs.lean b/Mathlib/Data/Nat/Prime/Defs.lean index c978c30544721..d062858f5ac79 100644 --- a/Mathlib/Data/Nat/Prime/Defs.lean +++ b/Mathlib/Data/Nat/Prime/Defs.lean @@ -4,8 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ import Batteries.Data.Nat.Gcd +import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Prime.Defs import Mathlib.Algebra.Ring.Nat +import Mathlib.Data.Nat.Sqrt import Mathlib.Order.Basic /-! diff --git a/Mathlib/Data/Nat/Size.lean b/Mathlib/Data/Nat/Size.lean index fd9311b427815..63929ccf4f5d8 100644 --- a/Mathlib/Data/Nat/Size.lean +++ b/Mathlib/Data/Nat/Size.lean @@ -3,6 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ +import Mathlib.Algebra.Group.Basic import Mathlib.Data.Nat.Bits /-! Lemmas about `size`. -/ diff --git a/Mathlib/Data/Set/Enumerate.lean b/Mathlib/Data/Set/Enumerate.lean index 908974989a9e3..c5858d0c5e375 100644 --- a/Mathlib/Data/Set/Enumerate.lean +++ b/Mathlib/Data/Set/Enumerate.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Algebra.Group.Basic -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic import Mathlib.Tactic.Common import Mathlib.Data.Set.Basic diff --git a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean index 3ce9f9ecba4d9..318e8a450260d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Quaternion.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Julian Kuelshammer -/ import Mathlib.Data.ZMod.Basic -import Mathlib.Algebra.Group.Nat import Mathlib.Tactic.IntervalCases import Mathlib.GroupTheory.SpecificGroups.Dihedral import Mathlib.GroupTheory.SpecificGroups.Cyclic diff --git a/Mathlib/Tactic/Sat/FromLRAT.lean b/Mathlib/Tactic/Sat/FromLRAT.lean index c1962b10c84a8..db7f18cb54a91 100644 --- a/Mathlib/Tactic/Sat/FromLRAT.lean +++ b/Mathlib/Tactic/Sat/FromLRAT.lean @@ -3,7 +3,8 @@ Copyright (c) 2022 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic +import Mathlib.Tactic.ByContra /-! # `lrat_proof` command diff --git a/MathlibTest/levenshtein.lean b/MathlibTest/levenshtein.lean index 5d2c2d34f0d48..53932d20b2d56 100644 --- a/MathlibTest/levenshtein.lean +++ b/MathlibTest/levenshtein.lean @@ -1,5 +1,5 @@ import Mathlib.Data.List.EditDistance.Defs -import Mathlib.Algebra.Group.Nat +import Mathlib.Algebra.Group.Nat.Basic #guard (suffixLevenshtein Levenshtein.defaultCost "kitten".toList "sitting".toList).1 = From 6017a6d9107cfaa4480f8b705a2d1d41b76003a4 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 08:36:26 +1100 Subject: [PATCH 041/250] shake --- Mathlib/Data/List/Indexes.lean | 1 - Mathlib/SetTheory/Cardinal/Basic.lean | 1 - 2 files changed, 2 deletions(-) diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 6c8ce5ff6c794..2097c5adf2cb5 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Mathlib.Order.Lattice import Mathlib.Data.List.Basic /-! diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index b309bcdad858b..2b76c69c48ddc 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -14,7 +14,6 @@ import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.InitialSeg import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.SetTheory.Cardinal.SchroederBernstein -import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Algebra.Order.Ring.Canonical From 9e7e42735e19836fc0b9bdccdedee5821afdc982 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Fri, 22 Nov 2024 22:03:23 +0000 Subject: [PATCH 042/250] chore(Algebra/Order/SuccPred/TypeTags): fix defeq abuse with Additive and Multiplicative (#19379) these simp lemmas now don't use defeq abuse between `A` and `Multiplicative A`. --- Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 7c69e6dc40162..6cda0bda7a3e5 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -36,15 +36,19 @@ namespace Order open Additive Multiplicative @[simp] lemma succ_ofMul [Preorder X] [SuccOrder X] (x : X) : succ (ofMul x) = ofMul (succ x) := rfl -@[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : X) : succ (toMul x) = toMul (succ x) := rfl +@[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : Additive X) : + succ (toMul x) = toMul (succ x) := rfl @[simp] lemma succ_ofAdd [Preorder X] [SuccOrder X] (x : X) : succ (ofAdd x) = ofAdd (succ x) := rfl -@[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : X) : succ (toAdd x) = toAdd (succ x) := rfl +@[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : Multiplicative X) : + succ (toAdd x) = toAdd (succ x) := rfl @[simp] lemma pred_ofMul [Preorder X] [PredOrder X] (x : X) : pred (ofMul x) = ofMul (pred x) := rfl -@[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : X) : pred (toMul x) = toMul (pred x) := rfl +@[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : Additive X) : + pred (toMul x) = toMul (pred x) := rfl @[simp] lemma pred_ofAdd [Preorder X] [PredOrder X] (x : X) : pred (ofAdd x) = ofAdd (pred x) := rfl -@[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : X) : pred (toAdd x) = toAdd (pred x) := rfl +@[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : Multiplicative X) : + pred (toAdd x) = toAdd (pred x) := rfl end Order From 65cc62903990b71383fddf0866f6a21db3826a24 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 22:03:24 +0000 Subject: [PATCH 043/250] chore: give the 'post or update summary comment' job a descriptive name (#19382) It can be useful to have a useful *job* (and not *workflow step*) name, e.g. in case one looks at JSON information in the queueboard data, about failed CI. --- .github/workflows/PR_summary.yml | 1 + 1 file changed, 1 insertion(+) diff --git a/.github/workflows/PR_summary.yml b/.github/workflows/PR_summary.yml index 09e8747fdd764..6b854a59292d0 100644 --- a/.github/workflows/PR_summary.yml +++ b/.github/workflows/PR_summary.yml @@ -5,6 +5,7 @@ on: jobs: build: + name: "post-or-update-summary-comment" runs-on: ubuntu-latest steps: From b0890aa4ad52262811ba72d7f467848c16b5a497 Mon Sep 17 00:00:00 2001 From: grunweg Date: Fri, 22 Nov 2024 22:03:26 +0000 Subject: [PATCH 044/250] chore: rename no_lints_prime_decls.txt to nolints_... (#19383) For consistency with the other nolints files. Suggested in #19275. --- Mathlib/Tactic/Linter/DocPrime.lean | 6 +++--- scripts/README.md | 3 +-- .../{no_lints_prime_decls.txt => nolints_prime_decls.txt} | 0 scripts/technical-debt-metrics.sh | 2 +- 4 files changed, 5 insertions(+), 6 deletions(-) rename scripts/{no_lints_prime_decls.txt => nolints_prime_decls.txt} (100%) diff --git a/Mathlib/Tactic/Linter/DocPrime.lean b/Mathlib/Tactic/Linter/DocPrime.lean index 7221a12e1ad7f..c2a8a38fbcdcf 100644 --- a/Mathlib/Tactic/Linter/DocPrime.lean +++ b/Mathlib/Tactic/Linter/DocPrime.lean @@ -26,7 +26,7 @@ namespace Mathlib.Linter The "docPrime" linter emits a warning on declarations that have no doc-string and whose name ends with a `'`. -The file `scripts/no_lints_prime_decls.txt` contains a list of temporary exceptions to this linter. +The file `scripts/nolints_prime_decls.txt` contains a list of temporary exceptions to this linter. This list should not be appended to, and become emptied over time. -/ register_option linter.docPrime : Bool := { @@ -63,8 +63,8 @@ def docPrimeLinter : Linter where run := withSetOptionIn fun stx ↦ do relative to the unprimed version, or an explanation as to why no better naming scheme \ is possible." if docstring[0][1].getAtomVal.isEmpty && declName.toString.back == '\'' then - if ← System.FilePath.pathExists "scripts/no_lints_prime_decls.txt" then - if (← IO.FS.lines "scripts/no_lints_prime_decls.txt").contains declName.toString then + if ← System.FilePath.pathExists "scripts/nolints_prime_decls.txt" then + if (← IO.FS.lines "scripts/nolints_prime_decls.txt").contains declName.toString then return else Linter.logLint linter.docPrime declId msg diff --git a/scripts/README.md b/scripts/README.md index 8f55ce30822bb..44156380d5e62 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -80,8 +80,7 @@ to learn about it as well! **Data files with linter exceptions** - `nolints.json` contains exceptions for all `env_linter`s in mathlib. For permanent and deliberate exceptions, add a `@[nolint lintername]` in the .lean file instead. -- `no_lints_prime_decls.txt` - contains temporary exceptions for the `docPrime` linter +- `nolints_prime_decls.txt` contains temporary exceptions for the `docPrime` linter Both of these files should tend to zero over time; please do not add new entries to these files. PRs removing (the need for) entries are welcome. diff --git a/scripts/no_lints_prime_decls.txt b/scripts/nolints_prime_decls.txt similarity index 100% rename from scripts/no_lints_prime_decls.txt rename to scripts/nolints_prime_decls.txt diff --git a/scripts/technical-debt-metrics.sh b/scripts/technical-debt-metrics.sh index eb24ded1ce153..b3f3df14ec9aa 100755 --- a/scripts/technical-debt-metrics.sh +++ b/scripts/technical-debt-metrics.sh @@ -113,7 +113,7 @@ printf '%s|%s\n' "$(grep -c 'docBlame' scripts/nolints.json)" "documentation nol printf '%s|%s\n' "$(git grep '^set_option linter.style.longFile [0-9]*' Mathlib | wc -l)" "large files" printf '%s|%s\n' "$(git grep "^open .*Classical" | grep -v " in$" -c)" "bare open (scoped) Classical" -printf '%s|%s\n' "$(wc -l < scripts/no_lints_prime_decls.txt)" "exceptions for the docPrime linter" +printf '%s|%s\n' "$(wc -l < scripts/nolints_prime_decls.txt)" "exceptions for the docPrime linter" deprecatedFiles="$(git ls-files '**/Deprecated/*.lean' | xargs wc -l | sed 's=^ *==')" From 4fe03b813c934955c13c0e74e9ba02d5e8449d06 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 00:05:11 +0000 Subject: [PATCH 045/250] chore(Algebra/Group/TypeTags): Remove porting notes related to lean4#1910, use new notation everywhere (#19369) part 1: make applied occurances of `Additive.toAdd` and `Multiplicative.toMul` use newly available dot notation. Co-authored-by: Eric Wieser --- Mathlib/Algebra/AddConstMap/Basic.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 2 +- .../Algebra/BigOperators/Group/Finset.lean | 12 ++-- Mathlib/Algebra/Group/Action/TypeTags.lean | 8 +-- Mathlib/Algebra/Group/AddChar.lean | 6 +- Mathlib/Algebra/Group/Aut.lean | 4 +- Mathlib/Algebra/Group/Equiv/TypeTags.lean | 20 +++--- Mathlib/Algebra/Group/Even.lean | 4 +- Mathlib/Algebra/Group/Int.lean | 4 +- Mathlib/Algebra/Group/Nat/TypeTags.lean | 2 +- .../Algebra/Group/Submonoid/Membership.lean | 2 +- Mathlib/Algebra/Group/TypeTags/Basic.lean | 65 +++++++++---------- Mathlib/Algebra/Group/TypeTags/Hom.lean | 12 ++-- Mathlib/Algebra/MonoidAlgebra/Basic.lean | 2 +- Mathlib/Algebra/MonoidAlgebra/Defs.lean | 4 +- Mathlib/Algebra/MonoidAlgebra/Division.lean | 6 +- Mathlib/Algebra/MonoidAlgebra/Grading.lean | 4 +- Mathlib/Algebra/Order/Archimedean/Basic.lean | 4 +- .../Order/GroupWithZero/Canonical.lean | 8 +-- Mathlib/Algebra/Order/Monoid/ToMulBot.lean | 2 +- .../Order/Monoid/Unbundled/TypeTags.lean | 8 +-- Mathlib/Algebra/Order/SuccPred/TypeTags.lean | 8 +-- .../Analysis/Normed/Group/Constructions.lean | 12 ++-- Mathlib/Data/Complex/Exponential.lean | 4 +- Mathlib/Data/Int/Cast/Lemmas.lean | 6 +- Mathlib/Data/Int/WithZero.lean | 8 +-- Mathlib/Data/Nat/Cast/Basic.lean | 6 +- Mathlib/Data/ZMod/IntUnitsPower.lean | 14 ++-- .../RotationNumber/TranslationNumber.lean | 2 +- .../GroupTheory/FiniteAbelian/Duality.lean | 4 +- .../AffineSpace/AffineEquiv.lean | 2 +- .../NumberField/Units/DirichletTheorem.lean | 4 +- Mathlib/RepresentationTheory/Basic.lean | 2 +- .../GroupCohomology/LowDegree.lean | 2 +- Mathlib/RepresentationTheory/Rep.lean | 2 +- Mathlib/RingTheory/FreeCommRing.lean | 2 +- Mathlib/RingTheory/LaurentSeries.lean | 2 +- Mathlib/Topology/Constructions.lean | 4 +- Mathlib/Topology/EMetricSpace/Defs.lean | 4 +- Mathlib/Topology/MetricSpace/Defs.lean | 8 +-- .../Topology/MetricSpace/IsometricSMul.lean | 12 ++-- 41 files changed, 142 insertions(+), 147 deletions(-) diff --git a/Mathlib/Algebra/AddConstMap/Basic.lean b/Mathlib/Algebra/AddConstMap/Basic.lean index b8a261ee2098e..25685a81e92e9 100644 --- a/Mathlib/Algebra/AddConstMap/Basic.lean +++ b/Mathlib/Algebra/AddConstMap/Basic.lean @@ -433,7 +433,7 @@ variable {G : Type*} [AddMonoid G] {a : G} as a monoid homomorphism from `Multiplicative G` to `G →+c[a, a] G`. -/ @[simps! (config := .asFn)] def addLeftHom : Multiplicative G →* (G →+c[a, a] G) where - toFun c := Multiplicative.toAdd c +ᵥ .id + toFun c := c.toAdd +ᵥ .id map_one' := by ext; apply zero_add map_mul' _ _ := by ext; apply add_assoc diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 2ab53a91f84e3..3f6183df46993 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -358,7 +358,7 @@ theorem constVAdd_add (v₁ v₂ : G) : constVAdd P (v₁ + v₂) = constVAdd P /-- `Equiv.constVAdd` as a homomorphism from `Multiplicative G` to `Equiv.perm P` -/ def constVAddHom : Multiplicative G →* Equiv.Perm P where - toFun v := constVAdd P (Multiplicative.toAdd v) + toFun v := constVAdd P (v.toAdd) map_one' := constVAdd_zero G P map_mul' := constVAdd_add P diff --git a/Mathlib/Algebra/BigOperators/Group/Finset.lean b/Mathlib/Algebra/BigOperators/Group/Finset.lean index a8d12379bbcce..282177e1ba7c0 100644 --- a/Mathlib/Algebra/BigOperators/Group/Finset.lean +++ b/Mathlib/Algebra/BigOperators/Group/Finset.lean @@ -2197,7 +2197,7 @@ variable [Monoid α] theorem ofMul_list_prod (s : List α) : ofMul s.prod = (s.map ofMul).sum := by simp [ofMul]; rfl @[simp] -theorem toMul_list_sum (s : List (Additive α)) : toMul s.sum = (s.map toMul).prod := by +theorem toMul_list_sum (s : List (Additive α)) : s.sum.toMul = (s.map toMul).prod := by simp [toMul, ofMul]; rfl end Monoid @@ -2210,7 +2210,7 @@ variable [AddMonoid α] theorem ofAdd_list_prod (s : List α) : ofAdd s.sum = (s.map ofAdd).prod := by simp [ofAdd]; rfl @[simp] -theorem toAdd_list_sum (s : List (Multiplicative α)) : toAdd s.prod = (s.map toAdd).sum := by +theorem toAdd_list_sum (s : List (Multiplicative α)) : s.prod.toAdd = (s.map toAdd).sum := by simp [toAdd, ofAdd]; rfl end AddMonoid @@ -2224,7 +2224,7 @@ theorem ofMul_multiset_prod (s : Multiset α) : ofMul s.prod = (s.map ofMul).sum simp [ofMul]; rfl @[simp] -theorem toMul_multiset_sum (s : Multiset (Additive α)) : toMul s.sum = (s.map toMul).prod := by +theorem toMul_multiset_sum (s : Multiset (Additive α)) : s.sum.toMul = (s.map toMul).prod := by simp [toMul, ofMul]; rfl @[simp] @@ -2233,7 +2233,7 @@ theorem ofMul_prod (s : Finset ι) (f : ι → α) : ofMul (∏ i ∈ s, f i) = @[simp] theorem toMul_sum (s : Finset ι) (f : ι → Additive α) : - toMul (∑ i ∈ s, f i) = ∏ i ∈ s, toMul (f i) := + (∑ i ∈ s, f i).toMul = ∏ i ∈ s, (f i).toMul := rfl end CommMonoid @@ -2248,7 +2248,7 @@ theorem ofAdd_multiset_prod (s : Multiset α) : ofAdd s.sum = (s.map ofAdd).prod @[simp] theorem toAdd_multiset_sum (s : Multiset (Multiplicative α)) : - toAdd s.prod = (s.map toAdd).sum := by + s.prod.toAdd = (s.map toAdd).sum := by simp [toAdd, ofAdd]; rfl @[simp] @@ -2257,7 +2257,7 @@ theorem ofAdd_sum (s : Finset ι) (f : ι → α) : ofAdd (∑ i ∈ s, f i) = @[simp] theorem toAdd_prod (s : Finset ι) (f : ι → Multiplicative α) : - toAdd (∏ i ∈ s, f i) = ∑ i ∈ s, toAdd (f i) := + (∏ i ∈ s, f i).toAdd = ∑ i ∈ s, (f i).toAdd := rfl end AddCommMonoid diff --git a/Mathlib/Algebra/Group/Action/TypeTags.lean b/Mathlib/Algebra/Group/Action/TypeTags.lean index c6ceefe332560..1e65ced1a6024 100644 --- a/Mathlib/Algebra/Group/Action/TypeTags.lean +++ b/Mathlib/Algebra/Group/Action/TypeTags.lean @@ -24,15 +24,15 @@ section open Additive Multiplicative -instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (toMul a • ·) +instance Additive.vadd [SMul α β] : VAdd (Additive α) β where vadd a := (a.toMul • ·) -instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (toAdd a +ᵥ ·) +instance Multiplicative.smul [VAdd α β] : SMul (Multiplicative α) β where smul a := (a.toAdd +ᵥ ·) -@[simp] lemma toMul_smul [SMul α β] (a) (b : β) : (toMul a : α) • b = a +ᵥ b := rfl +@[simp] lemma toMul_smul [SMul α β] (a:Additive α) (b : β) : (a.toMul : α) • b = a +ᵥ b := rfl @[simp] lemma ofMul_vadd [SMul α β] (a : α) (b : β) : ofMul a +ᵥ b = a • b := rfl -@[simp] lemma toAdd_vadd [VAdd α β] (a) (b : β) : (toAdd a : α) +ᵥ b = a • b := rfl +@[simp] lemma toAdd_vadd [VAdd α β] (a:Multiplicative α) (b : β) : (a.toAdd : α) +ᵥ b = a • b := rfl @[simp] lemma ofAdd_smul [VAdd α β] (a : α) (b : β) : ofAdd a • b = a +ᵥ b := rfl diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 23505cff4f752..4b0c0f3e8d9bd 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -115,7 +115,7 @@ def toMonoidHom (φ : AddChar A M) : Multiplicative A →* M where -- this instance was a bad idea and conflicted with `instFunLike` above @[simp] lemma toMonoidHom_apply (ψ : AddChar A M) (a : Multiplicative A) : - ψ.toMonoidHom a = ψ (Multiplicative.toAdd a) := + ψ.toMonoidHom a = ψ a.toAdd := rfl /-- An additive character maps multiples by natural numbers to powers. -/ @@ -142,7 +142,7 @@ def toMonoidHomEquiv : AddChar A M ≃ (Multiplicative A →* M) where ⇑(toMonoidHomEquiv.symm ψ) = ψ ∘ Multiplicative.ofAdd := rfl @[simp] lemma toMonoidHomEquiv_apply (ψ : AddChar A M) (a : Multiplicative A) : - toMonoidHomEquiv ψ a = ψ (Multiplicative.toAdd a) := rfl + toMonoidHomEquiv ψ a = ψ a.toAdd := rfl @[simp] lemma toMonoidHomEquiv_symm_apply (ψ : Multiplicative A →* M) (a : A) : toMonoidHomEquiv.symm ψ a = ψ (Multiplicative.ofAdd a) := rfl @@ -180,7 +180,7 @@ lemma coe_toAddMonoidHomEquiv (ψ : AddChar A M) : toAddMonoidHomEquiv ψ a = Additive.ofMul (ψ a) := rfl @[simp] lemma toAddMonoidHomEquiv_symm_apply (ψ : A →+ Additive M) (a : A) : - toAddMonoidHomEquiv.symm ψ a = Additive.toMul (ψ a) := rfl + toAddMonoidHomEquiv.symm ψ a = (ψ a).toMul := rfl /-- The trivial additive character (sending everything to `1`). -/ instance instOne : One (AddChar A M) := toMonoidHomEquiv.one diff --git a/Mathlib/Algebra/Group/Aut.lean b/Mathlib/Algebra/Group/Aut.lean index ac678be0258c9..012f362054268 100644 --- a/Mathlib/Algebra/Group/Aut.lean +++ b/Mathlib/Algebra/Group/Aut.lean @@ -260,9 +260,9 @@ theorem conj_symm_apply [AddGroup G] (g h : G) : (conj g).symm h = -g + h + g := -- Porting note: the exact translation of this mathlib3 lemma would be`(-conj g) h = -g + h + g`, -- but this no longer pass the simp_nf linter, as the LHS simplifies by `toMul_neg` to --- `(Additive.toMul (conj g))⁻¹`. +-- `(conj g).toMul⁻¹`. @[simp] -theorem conj_inv_apply [AddGroup G] (g h : G) : (Additive.toMul (conj g))⁻¹ h = -g + h + g := +theorem conj_inv_apply [AddGroup G] (g h : G) : (conj g).toMul⁻¹ h = -g + h + g := rfl /-- Isomorphic additive groups have isomorphic automorphism groups. -/ diff --git a/Mathlib/Algebra/Group/Equiv/TypeTags.lean b/Mathlib/Algebra/Group/Equiv/TypeTags.lean index f8e14c0ceb128..12338d2419e5a 100644 --- a/Mathlib/Algebra/Group/Equiv/TypeTags.lean +++ b/Mathlib/Algebra/Group/Equiv/TypeTags.lean @@ -118,8 +118,8 @@ and multiplicative endomorphisms of `Multiplicative A`. -/ @[simps] def MulEquiv.piMultiplicative (K : ι → Type*) [∀ i, Add (K i)] : Multiplicative (∀ i : ι, K i) ≃* (∀ i : ι, Multiplicative (K i)) where - toFun x := fun i ↦ Multiplicative.ofAdd <| Multiplicative.toAdd x i - invFun x := Multiplicative.ofAdd fun i ↦ Multiplicative.toAdd (x i) + toFun x := fun i ↦ Multiplicative.ofAdd <| x.toAdd i + invFun x := Multiplicative.ofAdd fun i ↦ (x i).toAdd left_inv _ := rfl right_inv _ := rfl map_mul' _ _ := rfl @@ -134,8 +134,8 @@ abbrev MulEquiv.funMultiplicative [Add G] : @[simps] def AddEquiv.piAdditive (K : ι → Type*) [∀ i, Mul (K i)] : Additive (∀ i : ι, K i) ≃+ (∀ i : ι, Additive (K i)) where - toFun x := fun i ↦ Additive.ofMul <| Additive.toMul x i - invFun x := Additive.ofMul fun i ↦ Additive.toMul (x i) + toFun x := fun i ↦ Additive.ofMul <| x.toMul i + invFun x := Additive.ofMul fun i ↦ (x i).toMul left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl @@ -164,9 +164,9 @@ def MulEquiv.multiplicativeAdditive [MulOneClass H] : Multiplicative (Additive H @[simps] def MulEquiv.prodMultiplicative [Add G] [Add H] : Multiplicative (G × H) ≃* Multiplicative G × Multiplicative H where - toFun x := (Multiplicative.ofAdd (Multiplicative.toAdd x).1, - Multiplicative.ofAdd (Multiplicative.toAdd x).2) - invFun := fun (x, y) ↦ Multiplicative.ofAdd (Multiplicative.toAdd x, Multiplicative.toAdd y) + toFun x := (Multiplicative.ofAdd x.toAdd.1, + Multiplicative.ofAdd x.toAdd.2) + invFun := fun (x, y) ↦ Multiplicative.ofAdd (x.toAdd, y.toAdd) left_inv _ := rfl right_inv _ := rfl map_mul' _ _ := rfl @@ -175,9 +175,9 @@ def MulEquiv.prodMultiplicative [Add G] [Add H] : @[simps] def AddEquiv.prodAdditive [Mul G] [Mul H] : Additive (G × H) ≃+ Additive G × Additive H where - toFun x := (Additive.ofMul (Additive.toMul x).1, - Additive.ofMul (Additive.toMul x).2) - invFun := fun (x, y) ↦ Additive.ofMul (Additive.toMul x, Additive.toMul y) + toFun x := (Additive.ofMul x.toMul.1, + Additive.ofMul x.toMul.2) + invFun := fun (x, y) ↦ Additive.ofMul (x.toMul, y.toMul) left_inv _ := rfl right_inv _ := rfl map_add' _ _ := rfl diff --git a/Mathlib/Algebra/Group/Even.lean b/Mathlib/Algebra/Group/Even.lean index febd4c5c50c99..8df0d5b3613c9 100644 --- a/Mathlib/Algebra/Group/Even.lean +++ b/Mathlib/Algebra/Group/Even.lean @@ -62,7 +62,7 @@ instance [DecidablePred (IsSquare : α → Prop)] : DecidablePred (IsSquare : α lemma even_ofMul_iff {a : α} : Even (Additive.ofMul a) ↔ IsSquare a := Iff.rfl @[simp] -lemma isSquare_toMul_iff {a : Additive α} : IsSquare (Additive.toMul a) ↔ Even a := Iff.rfl +lemma isSquare_toMul_iff {a : Additive α} : IsSquare (a.toMul) ↔ Even a := Iff.rfl instance Additive.instDecidablePredEven [DecidablePred (IsSquare : α → Prop)] : DecidablePred (Even : Additive α → Prop) := @@ -76,7 +76,7 @@ variable [Add α] @[simp] lemma isSquare_ofAdd_iff {a : α} : IsSquare (Multiplicative.ofAdd a) ↔ Even a := Iff.rfl @[simp] -lemma even_toAdd_iff {a : Multiplicative α} : Even (Multiplicative.toAdd a) ↔ IsSquare a := Iff.rfl +lemma even_toAdd_iff {a : Multiplicative α} : Even a.toAdd ↔ IsSquare a := Iff.rfl instance Multiplicative.instDecidablePredIsSquare [DecidablePred (Even : α → Prop)] : DecidablePred (IsSquare : Multiplicative α → Prop) := diff --git a/Mathlib/Algebra/Group/Int.lean b/Mathlib/Algebra/Group/Int.lean index 948531494a03b..e9b98108c5110 100644 --- a/Mathlib/Algebra/Group/Int.lean +++ b/Mathlib/Algebra/Group/Int.lean @@ -74,9 +74,9 @@ section Multiplicative open Multiplicative -lemma toAdd_pow (a : Multiplicative ℤ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_pow (a : Multiplicative ℤ) (b : ℕ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ -lemma toAdd_zpow (a : Multiplicative ℤ) (b : ℤ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_zpow (a : Multiplicative ℤ) (b : ℤ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ @[simp] lemma ofAdd_mul (a b : ℤ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_zpow ..).symm diff --git a/Mathlib/Algebra/Group/Nat/TypeTags.lean b/Mathlib/Algebra/Group/Nat/TypeTags.lean index 0efe1e8dc736b..a312f69a4aab7 100644 --- a/Mathlib/Algebra/Group/Nat/TypeTags.lean +++ b/Mathlib/Algebra/Group/Nat/TypeTags.lean @@ -17,7 +17,7 @@ open Multiplicative namespace Nat -lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : toAdd (a ^ b) = toAdd a * b := mul_comm _ _ +lemma toAdd_pow (a : Multiplicative ℕ) (b : ℕ) : (a ^ b).toAdd = a.toAdd * b := mul_comm _ _ @[simp] lemma ofAdd_mul (a b : ℕ) : ofAdd (a * b) = ofAdd a ^ b := (toAdd_pow _ _).symm diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index cc9d5be5714ae..2f07619bf80ce 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -475,7 +475,7 @@ when it is injective. The inverse is given by the logarithms. -/ @[simps] def powLogEquiv [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n ^ m) : Multiplicative ℕ ≃* powers n where - toFun m := pow n (Multiplicative.toAdd m) + toFun m := pow n m.toAdd invFun m := Multiplicative.ofAdd (log m) left_inv := log_pow_eq_self h right_inv := pow_log_eq_self diff --git a/Mathlib/Algebra/Group/TypeTags/Basic.lean b/Mathlib/Algebra/Group/TypeTags/Basic.lean index c506dc1cf3541..152313ad2ebc6 100644 --- a/Mathlib/Algebra/Group/TypeTags/Basic.lean +++ b/Mathlib/Algebra/Group/TypeTags/Basic.lean @@ -27,11 +27,6 @@ We also define instances `Additive.*` and `Multiplicative.*` that actually trans This file is similar to `Order.Synonym`. -## Porting notes - -- Since bundled morphism applications that rely on `CoeFun` currently don't work, they are ported - as `toFoo a` rather than `a.toFoo` for now. (https://github.com/leanprover/lean4/issues/1910) - -/ assert_not_exists MonoidWithZero @@ -77,7 +72,7 @@ protected lemma «exists» {p : Additive α → Prop} : (∃ a, p a) ↔ ∃ a, /-- Recursion principle for `Additive`, supported by `cases` and `induction`. -/ @[elab_as_elim, cases_eliminator, induction_eliminator] def rec {motive : Additive α → Sort*} (ofMul : ∀ a, motive (ofMul a)) : ∀ a, motive a := - fun a => ofMul (toMul a) + fun a => ofMul (a.toMul) end Additive @@ -107,7 +102,7 @@ protected lemma «exists» {p : Multiplicative α → Prop} : (∃ a, p a) ↔ /-- Recursion principle for `Multiplicative`, supported by `cases` and `induction`. -/ @[elab_as_elim, cases_eliminator, induction_eliminator] def rec {motive : Multiplicative α → Sort*} (ofAdd : ∀ a, motive (ofAdd a)) : ∀ a, motive a := - fun a => ofAdd (toAdd a) + fun a => ofAdd (a.toAdd) end Multiplicative @@ -115,19 +110,19 @@ open Additive (ofMul toMul) open Multiplicative (ofAdd toAdd) @[simp] -theorem toAdd_ofAdd (x : α) : toAdd (ofAdd x) = x := +theorem toAdd_ofAdd (x : α) : (ofAdd x).toAdd = x := rfl @[simp] -theorem ofAdd_toAdd (x : Multiplicative α) : ofAdd (toAdd x) = x := +theorem ofAdd_toAdd (x : Multiplicative α) : ofAdd x.toAdd = x := rfl @[simp] -theorem toMul_ofMul (x : α) : toMul (ofMul x) = x := +theorem toMul_ofMul (x : α) : (ofMul x).toMul = x := rfl @[simp] -theorem ofMul_toMul (x : Additive α) : ofMul (toMul x) = x := +theorem ofMul_toMul (x : Additive α) : ofMul x.toMul = x := rfl instance [Subsingleton α] : Subsingleton (Additive α) := toMul.injective.subsingleton @@ -153,22 +148,22 @@ instance Multiplicative.instNontrivial [Nontrivial α] : Nontrivial (Multiplicat ofAdd.injective.nontrivial instance Additive.add [Mul α] : Add (Additive α) where - add x y := ofMul (toMul x * toMul y) + add x y := ofMul (x.toMul * y.toMul) instance Multiplicative.mul [Add α] : Mul (Multiplicative α) where - mul x y := ofAdd (toAdd x + toAdd y) + mul x y := ofAdd (x.toAdd + y.toAdd) @[simp] theorem ofAdd_add [Add α] (x y : α) : ofAdd (x + y) = ofAdd x * ofAdd y := rfl @[simp] -theorem toAdd_mul [Add α] (x y : Multiplicative α) : toAdd (x * y) = toAdd x + toAdd y := rfl +theorem toAdd_mul [Add α] (x y : Multiplicative α) : (x * y).toAdd = x.toAdd + y.toAdd := rfl @[simp] theorem ofMul_mul [Mul α] (x y : α) : ofMul (x * y) = ofMul x + ofMul y := rfl @[simp] -theorem toMul_add [Mul α] (x y : Additive α) : toMul (x + y) = toMul x * toMul y := rfl +theorem toMul_add [Mul α] (x y : Additive α) : (x + y).toMul = x.toMul * y.toMul := rfl instance Additive.addSemigroup [Semigroup α] : AddSemigroup (Additive α) := { Additive.add with add_assoc := @mul_assoc α _ } @@ -228,11 +223,11 @@ theorem ofMul_one [One α] : @Additive.ofMul α 1 = 0 := rfl theorem ofMul_eq_zero {A : Type*} [One A] {x : A} : Additive.ofMul x = 0 ↔ x = 1 := Iff.rfl @[simp] -theorem toMul_zero [One α] : toMul (0 : Additive α) = 1 := rfl +theorem toMul_zero [One α] : (0 : Additive α).toMul = 1 := rfl @[simp] lemma toMul_eq_one {α : Type*} [One α] {x : Additive α} : - Additive.toMul x = 1 ↔ x = 0 := + x.toMul = 1 ↔ x = 0 := Iff.rfl instance [Zero α] : One (Multiplicative α) := @@ -247,12 +242,12 @@ theorem ofAdd_eq_one {A : Type*} [Zero A] {x : A} : Multiplicative.ofAdd x = 1 Iff.rfl @[simp] -theorem toAdd_one [Zero α] : toAdd (1 : Multiplicative α) = 0 := +theorem toAdd_one [Zero α] : (1 : Multiplicative α).toAdd = 0 := rfl @[simp] lemma toAdd_eq_zero {α : Type*} [Zero α] {x : Multiplicative α} : - Multiplicative.toAdd x = 0 ↔ x = 1 := + x.toAdd = 0 ↔ x = 1 := Iff.rfl instance Additive.addZeroClass [MulOneClass α] : AddZeroClass (Additive α) where @@ -284,7 +279,7 @@ theorem ofMul_pow [Monoid α] (n : ℕ) (a : α) : ofMul (a ^ n) = n • ofMul a rfl @[simp] -theorem toMul_nsmul [Monoid α] (n : ℕ) (a : Additive α) : toMul (n • a) = toMul a ^ n := +theorem toMul_nsmul [Monoid α] (n : ℕ) (a : Additive α) : (n • a).toMul = a.toMul ^ n := rfl @[simp] @@ -292,7 +287,7 @@ theorem ofAdd_nsmul [AddMonoid α] (n : ℕ) (a : α) : ofAdd (n • a) = ofAdd rfl @[simp] -theorem toAdd_pow [AddMonoid α] (a : Multiplicative α) (n : ℕ) : toAdd (a ^ n) = n • toAdd a := +theorem toAdd_pow [AddMonoid α] (a : Multiplicative α) (n : ℕ) : (a ^ n).toAdd = n • a.toAdd := rfl instance Additive.addLeftCancelMonoid [LeftCancelMonoid α] : AddLeftCancelMonoid (Additive α) := @@ -316,39 +311,39 @@ instance Multiplicative.commMonoid [AddCommMonoid α] : CommMonoid (Multiplicati { Multiplicative.monoid, Multiplicative.commSemigroup with } instance Additive.neg [Inv α] : Neg (Additive α) := - ⟨fun x => ofAdd (toMul x)⁻¹⟩ + ⟨fun x => ofAdd x.toMul⁻¹⟩ @[simp] theorem ofMul_inv [Inv α] (x : α) : ofMul x⁻¹ = -ofMul x := rfl @[simp] -theorem toMul_neg [Inv α] (x : Additive α) : toMul (-x) = (toMul x)⁻¹ := +theorem toMul_neg [Inv α] (x : Additive α) : (-x).toMul = x.toMul⁻¹ := rfl instance Multiplicative.inv [Neg α] : Inv (Multiplicative α) := - ⟨fun x => ofMul (-toAdd x)⟩ + ⟨fun x => ofMul (-x.toAdd)⟩ @[simp] theorem ofAdd_neg [Neg α] (x : α) : ofAdd (-x) = (ofAdd x)⁻¹ := rfl @[simp] -theorem toAdd_inv [Neg α] (x : Multiplicative α) : toAdd x⁻¹ = -(toAdd x) := +theorem toAdd_inv [Neg α] (x : Multiplicative α) : x⁻¹.toAdd = -x.toAdd := rfl instance Additive.sub [Div α] : Sub (Additive α) where - sub x y := ofMul (toMul x / toMul y) + sub x y := ofMul (x.toMul / y.toMul) instance Multiplicative.div [Sub α] : Div (Multiplicative α) where - div x y := ofAdd (toAdd x - toAdd y) + div x y := ofAdd (x.toAdd - y.toAdd) @[simp] theorem ofAdd_sub [Sub α] (x y : α) : ofAdd (x - y) = ofAdd x / ofAdd y := rfl @[simp] -theorem toAdd_div [Sub α] (x y : Multiplicative α) : toAdd (x / y) = toAdd x - toAdd y := +theorem toAdd_div [Sub α] (x y : Multiplicative α) : (x / y).toAdd = x.toAdd - y.toAdd := rfl @[simp] @@ -356,7 +351,7 @@ theorem ofMul_div [Div α] (x y : α) : ofMul (x / y) = ofMul x - ofMul y := rfl @[simp] -theorem toMul_sub [Div α] (x y : Additive α) : toMul (x - y) = toMul x / toMul y := +theorem toMul_sub [Div α] (x y : Additive α) : (x - y).toMul = x.toMul / y.toMul := rfl instance Additive.involutiveNeg [InvolutiveInv α] : InvolutiveNeg (Additive α) := @@ -386,7 +381,7 @@ theorem ofMul_zpow [DivInvMonoid α] (z : ℤ) (a : α) : ofMul (a ^ z) = z • rfl @[simp] -theorem toMul_zsmul [DivInvMonoid α] (z : ℤ) (a : Additive α) : toMul (z • a) = toMul a ^ z := +theorem toMul_zsmul [DivInvMonoid α] (z : ℤ) (a : Additive α) : (z • a).toMul = a.toMul ^ z := rfl @[simp] @@ -394,7 +389,7 @@ theorem ofAdd_zsmul [SubNegMonoid α] (z : ℤ) (a : α) : ofAdd (z • a) = ofA rfl @[simp] -theorem toAdd_zpow [SubNegMonoid α] (a : Multiplicative α) (z : ℤ) : toAdd (a ^ z) = z • toAdd a := +theorem toAdd_zpow [SubNegMonoid α] (a : Multiplicative α) (z : ℤ) : (a ^ z).toAdd = z • a.toAdd := rfl instance Additive.subtractionMonoid [DivisionMonoid α] : SubtractionMonoid (Additive α) := @@ -434,8 +429,8 @@ This allows `Additive` to be used on bundled function types with a multiplicativ is often used for composition, without affecting the behavior of the function itself. -/ instance Additive.coeToFun {α : Type*} {β : α → Sort*} [CoeFun α β] : - CoeFun (Additive α) fun a => β (toMul a) := - ⟨fun a => CoeFun.coe (toMul a)⟩ + CoeFun (Additive α) fun a => β a.toMul := + ⟨fun a => CoeFun.coe a.toMul⟩ /-- If `α` has some additive structure and coerces to a function, then `Multiplicative α` should also coerce to the same function. @@ -444,8 +439,8 @@ This allows `Multiplicative` to be used on bundled function types with an additi is often used for composition, without affecting the behavior of the function itself. -/ instance Multiplicative.coeToFun {α : Type*} {β : α → Sort*} [CoeFun α β] : - CoeFun (Multiplicative α) fun a => β (toAdd a) := - ⟨fun a => CoeFun.coe (toAdd a)⟩ + CoeFun (Multiplicative α) fun a => β a.toAdd := + ⟨fun a => CoeFun.coe a.toAdd⟩ lemma Pi.mulSingle_multiplicativeOfAdd_eq {ι : Type*} [DecidableEq ι] {M : ι → Type*} [(i : ι) → AddMonoid (M i)] (i : ι) (a : M i) (j : ι) : diff --git a/Mathlib/Algebra/Group/TypeTags/Hom.lean b/Mathlib/Algebra/Group/TypeTags/Hom.lean index 749c54aa7a8b7..909615b47da55 100644 --- a/Mathlib/Algebra/Group/TypeTags/Hom.lean +++ b/Mathlib/Algebra/Group/TypeTags/Hom.lean @@ -23,12 +23,12 @@ open Multiplicative (ofAdd toAdd) def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : (α →+ β) ≃ (Multiplicative α →* Multiplicative β) where toFun f := { - toFun := fun a => ofAdd (f (toAdd a)) + toFun := fun a => ofAdd (f a.toAdd) map_mul' := f.map_add map_one' := f.map_zero } invFun f := { - toFun := fun a => toAdd (f (ofAdd a)) + toFun := fun a => f (ofAdd a) |>.toAdd map_add' := f.map_mul map_zero' := f.map_one } @@ -44,12 +44,12 @@ lemma AddMonoidHom.coe_toMultiplicative [AddZeroClass α] [AddZeroClass β] (f : def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : (α →* β) ≃ (Additive α →+ Additive β) where toFun f := { - toFun := fun a => ofMul (f (toMul a)) + toFun := fun a => ofMul (f a.toMul) map_add' := f.map_mul map_zero' := f.map_one } invFun f := { - toFun := fun a => toMul (f (ofMul a)) + toFun := fun a => (f (ofMul a)).toMul map_mul' := f.map_add map_one' := f.map_zero } @@ -70,7 +70,7 @@ def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : map_one' := f.map_zero } invFun f := { - toFun := fun a => toAdd (f (toMul a)) + toFun := fun a => (f a.toMul).toAdd map_add' := f.map_mul map_zero' := f.map_one } @@ -96,7 +96,7 @@ lemma MonoidHom.coe_toAdditive' [MulOneClass α] [AddZeroClass β] (f : α →* def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : (α →+ Additive β) ≃ (Multiplicative α →* β) where toFun f := { - toFun := fun a => toMul (f (toAdd a)) + toFun := fun a => (f a.toAdd).toMul map_mul' := f.map_add map_one' := f.map_zero } diff --git a/Mathlib/Algebra/MonoidAlgebra/Basic.lean b/Mathlib/Algebra/MonoidAlgebra/Basic.lean index 07f16b5034f0f..1e0a8e01e33eb 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Basic.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Basic.lean @@ -463,7 +463,7 @@ theorem lift_def (F : Multiplicative G →* A) : @[simp] theorem lift_symm_apply (F : k[G] →ₐ[k] A) (x : Multiplicative G) : - (lift k G A).symm F x = F (single (Multiplicative.toAdd x) 1) := + (lift k G A).symm F x = F (single x.toAdd 1) := rfl theorem lift_of (F : Multiplicative G →* A) (x : Multiplicative G) : diff --git a/Mathlib/Algebra/MonoidAlgebra/Defs.lean b/Mathlib/Algebra/MonoidAlgebra/Defs.lean index 93e53c6c0603b..41e3c6d67e134 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Defs.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Defs.lean @@ -1300,7 +1300,7 @@ end @[simp] theorem of_apply [AddZeroClass G] (a : Multiplicative G) : - of k G a = single (Multiplicative.toAdd a) 1 := + of k G a = single a.toAdd 1 := rfl @[simp] @@ -1324,7 +1324,7 @@ Note the order of the elements of the product are reversed compared to the argum -/ @[simps] def singleHom [AddZeroClass G] : k × Multiplicative G →* k[G] where - toFun a := single (Multiplicative.toAdd a.2) a.1 + toFun a := single a.2.toAdd a.1 map_one' := rfl map_mul' _a _b := single_mul_single.symm diff --git a/Mathlib/Algebra/MonoidAlgebra/Division.lean b/Mathlib/Algebra/MonoidAlgebra/Division.lean index e0f57c1d1a2b2..7b63bd6d1bbd6 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Division.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Division.lean @@ -83,13 +83,13 @@ theorem divOf_add (x : k[G]) (a b : G) : x /ᵒᶠ (a + b) = x /ᵒᶠ a /ᵒᶠ @[simps] noncomputable def divOfHom : Multiplicative G →* AddMonoid.End k[G] where toFun g := - { toFun := fun x => divOf x (Multiplicative.toAdd g) + { toFun := fun x => divOf x g.toAdd map_zero' := zero_divOf _ - map_add' := fun x y => add_divOf x y (Multiplicative.toAdd g) } + map_add' := fun x y => add_divOf x y g.toAdd } map_one' := AddMonoidHom.ext divOf_zero map_mul' g₁ g₂ := AddMonoidHom.ext fun _x => - (congr_arg _ (add_comm (Multiplicative.toAdd g₁) (Multiplicative.toAdd g₂))).trans + (congr_arg _ (add_comm g₁.toAdd g₂.toAdd)).trans (divOf_add _ _ _) theorem of'_mul_divOf (a : G) (x : k[G]) : of' k G a * x /ᵒᶠ a = x := by diff --git a/Mathlib/Algebra/MonoidAlgebra/Grading.lean b/Mathlib/Algebra/MonoidAlgebra/Grading.lean index 39f3e606eae43..66092e7d7d662 100644 --- a/Mathlib/Algebra/MonoidAlgebra/Grading.lean +++ b/Mathlib/Algebra/MonoidAlgebra/Grading.lean @@ -110,8 +110,8 @@ variable [AddMonoid M] [DecidableEq ι] [AddMonoid ι] [CommSemiring R] (f : M def decomposeAux : R[M] →ₐ[R] ⨁ i : ι, gradeBy R f i := AddMonoidAlgebra.lift R M _ { toFun := fun m => - DirectSum.of (fun i : ι => gradeBy R f i) (f (Multiplicative.toAdd m)) - ⟨Finsupp.single (Multiplicative.toAdd m) 1, single_mem_gradeBy _ _ _⟩ + DirectSum.of (fun i : ι => gradeBy R f i) (f m.toAdd) + ⟨Finsupp.single m.toAdd 1, single_mem_gradeBy _ _ _⟩ map_one' := DirectSum.of_eq_of_gradedMonoid_eq (by congr 2 <;> simp) diff --git a/Mathlib/Algebra/Order/Archimedean/Basic.lean b/Mathlib/Algebra/Order/Archimedean/Basic.lean index c2c50a58878f6..bf952ac20407b 100644 --- a/Mathlib/Algebra/Order/Archimedean/Basic.lean +++ b/Mathlib/Algebra/Order/Archimedean/Basic.lean @@ -64,11 +64,11 @@ instance OrderDual.instMulArchimedean [OrderedCommGroup α] [MulArchimedean α] instance Additive.instArchimedean [OrderedCommGroup α] [MulArchimedean α] : Archimedean (Additive α) := - ⟨fun x _ hy ↦ MulArchimedean.arch (toMul x) hy⟩ + ⟨fun x _ hy ↦ MulArchimedean.arch x.toMul hy⟩ instance Multiplicative.instMulArchimedean [OrderedAddCommGroup α] [Archimedean α] : MulArchimedean (Multiplicative α) := - ⟨fun x _ hy ↦ Archimedean.arch (toAdd x) hy⟩ + ⟨fun x _ hy ↦ Archimedean.arch x.toAdd hy⟩ variable {M : Type*} diff --git a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean index 95813028bae02..d26601f412fd7 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Canonical.lean @@ -85,7 +85,7 @@ instance instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual : LinearOrderedAddCommMonoidWithTop (Additive αᵒᵈ) := { Additive.orderedAddCommMonoid, Additive.linearOrder with top := (0 : α) - top_add' := fun a ↦ zero_mul (Additive.toMul a) + top_add' := fun a ↦ zero_mul a.toMul le_top := fun _ ↦ zero_le' } variable [NoZeroDivisors α] @@ -195,7 +195,7 @@ instance : LinearOrderedAddCommGroupWithTop (Additive αᵒᵈ) := { Additive.subNegMonoid, instLinearOrderedAddCommMonoidWithTopAdditiveOrderDual, Additive.instNontrivial with neg_top := inv_zero (G₀ := α) - add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : Additive.toMul a ≠ 0) } + add_neg_cancel := fun a ha ↦ mul_inv_cancel₀ (G₀ := α) (id ha : a.toMul ≠ 0) } @[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 @@ -219,7 +219,7 @@ theorem ofAdd_toDual_eq_zero_iff [LinearOrderedAddCommMonoidWithTop α] @[simp] theorem ofDual_toAdd_eq_top_iff [LinearOrderedAddCommMonoidWithTop α] - (x : Multiplicative αᵒᵈ) : OrderDual.ofDual (Multiplicative.toAdd x) = ⊤ ↔ x = 0 := Iff.rfl + (x : Multiplicative αᵒᵈ) : OrderDual.ofDual x.toAdd = ⊤ ↔ x = 0 := Iff.rfl @[simp] theorem ofAdd_bot [LinearOrderedAddCommMonoidWithTop α] : @@ -227,7 +227,7 @@ theorem ofAdd_bot [LinearOrderedAddCommMonoidWithTop α] : @[simp] theorem ofDual_toAdd_zero [LinearOrderedAddCommMonoidWithTop α] : - OrderDual.ofDual (Multiplicative.toAdd (0 : Multiplicative αᵒᵈ)) = ⊤ := rfl + OrderDual.ofDual (0 : Multiplicative αᵒᵈ).toAdd = ⊤ := rfl instance [LinearOrderedAddCommGroupWithTop α] : LinearOrderedCommGroupWithZero (Multiplicative αᵒᵈ) := diff --git a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean index ab578ee5f849e..79f4fd7e4a107 100644 --- a/Mathlib/Algebra/Order/Monoid/ToMulBot.lean +++ b/Mathlib/Algebra/Order/Monoid/ToMulBot.lean @@ -33,7 +33,7 @@ theorem toMulBot_zero : toMulBot (0 : WithZero (Multiplicative α)) = Multiplica @[simp] theorem toMulBot_coe (x : Multiplicative α) : - toMulBot ↑x = Multiplicative.ofAdd (↑(Multiplicative.toAdd x) : WithBot α) := + toMulBot ↑x = Multiplicative.ofAdd (↑x.toAdd : WithBot α) := rfl @[simp] diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index a975f6b4f3384..8a7c336d6901c 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -79,11 +79,11 @@ theorem ofMul_lt {a b : α} : ofMul a < ofMul b ↔ a < b := Iff.rfl @[simp] -theorem toMul_le {a b : Additive α} : toMul a ≤ toMul b ↔ a ≤ b := +theorem toMul_le {a b : Additive α} : a.toMul ≤ b.toMul ↔ a ≤ b := Iff.rfl @[simp] -theorem toMul_lt {a b : Additive α} : toMul a < toMul b ↔ a < b := +theorem toMul_lt {a b : Additive α} : a.toMul < b.toMul ↔ a < b := Iff.rfl end Additive @@ -101,11 +101,11 @@ theorem ofAdd_lt {a b : α} : ofAdd a < ofAdd b ↔ a < b := Iff.rfl @[simp] -theorem toAdd_le {a b : Multiplicative α} : toAdd a ≤ toAdd b ↔ a ≤ b := +theorem toAdd_le {a b : Multiplicative α} : a.toAdd ≤ b.toAdd ↔ a ≤ b := Iff.rfl @[simp] -theorem toAdd_lt {a b : Multiplicative α} : toAdd a < toAdd b ↔ a < b := +theorem toAdd_lt {a b : Multiplicative α} : a.toAdd < b.toAdd ↔ a < b := Iff.rfl end Multiplicative diff --git a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean index 6cda0bda7a3e5..4ee0095a26d7c 100644 --- a/Mathlib/Algebra/Order/SuccPred/TypeTags.lean +++ b/Mathlib/Algebra/Order/SuccPred/TypeTags.lean @@ -37,18 +37,18 @@ open Additive Multiplicative @[simp] lemma succ_ofMul [Preorder X] [SuccOrder X] (x : X) : succ (ofMul x) = ofMul (succ x) := rfl @[simp] lemma succ_toMul [Preorder X] [SuccOrder X] (x : Additive X) : - succ (toMul x) = toMul (succ x) := rfl + succ x.toMul = (succ x).toMul := rfl @[simp] lemma succ_ofAdd [Preorder X] [SuccOrder X] (x : X) : succ (ofAdd x) = ofAdd (succ x) := rfl @[simp] lemma succ_toAdd [Preorder X] [SuccOrder X] (x : Multiplicative X) : - succ (toAdd x) = toAdd (succ x) := rfl + succ x.toAdd = (succ x).toAdd := rfl @[simp] lemma pred_ofMul [Preorder X] [PredOrder X] (x : X) : pred (ofMul x) = ofMul (pred x) := rfl @[simp] lemma pred_toMul [Preorder X] [PredOrder X] (x : Additive X) : - pred (toMul x) = toMul (pred x) := rfl + pred x.toMul = (pred x).toMul := rfl @[simp] lemma pred_ofAdd [Preorder X] [PredOrder X] (x : X) : pred (ofAdd x) = ofAdd (pred x) := rfl @[simp] lemma pred_toAdd [Preorder X] [PredOrder X] (x : Multiplicative X) : - pred (toAdd x) = toAdd (pred x) := rfl + pred x.toAdd = (pred x).toAdd := rfl end Order diff --git a/Mathlib/Analysis/Normed/Group/Constructions.lean b/Mathlib/Analysis/Normed/Group/Constructions.lean index c9ee9cc5d178f..95eb19382e965 100644 --- a/Mathlib/Analysis/Normed/Group/Constructions.lean +++ b/Mathlib/Analysis/Normed/Group/Constructions.lean @@ -103,11 +103,11 @@ variable [Norm E] instance Additive.toNorm : Norm (Additive E) := ‹Norm E› instance Multiplicative.toNorm : Norm (Multiplicative E) := ‹Norm E› -@[simp] lemma norm_toMul (x) : ‖(toMul x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_toMul (x : Additive E) : ‖(x.toMul : E)‖ = ‖x‖ := rfl @[simp] lemma norm_ofMul (x : E) : ‖ofMul x‖ = ‖x‖ := rfl -@[simp] lemma norm_toAdd (x) : ‖(toAdd x : E)‖ = ‖x‖ := rfl +@[simp] lemma norm_toAdd (x : Multiplicative E) : ‖(x.toAdd : E)‖ = ‖x‖ := rfl @[simp] lemma norm_ofAdd (x : E) : ‖ofAdd x‖ = ‖x‖ := rfl @@ -120,23 +120,23 @@ instance Additive.toNNNorm : NNNorm (Additive E) := ‹NNNorm E› instance Multiplicative.toNNNorm : NNNorm (Multiplicative E) := ‹NNNorm E› -@[simp] lemma nnnorm_toMul (x) : ‖(toMul x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_toMul (x : Additive E) : ‖(x.toMul : E)‖₊ = ‖x‖₊ := rfl @[simp] lemma nnnorm_ofMul (x : E) : ‖ofMul x‖₊ = ‖x‖₊ := rfl -@[simp] lemma nnnorm_toAdd (x) : ‖(toAdd x : E)‖₊ = ‖x‖₊ := rfl +@[simp] lemma nnnorm_toAdd (x : Multiplicative E) : ‖(x.toAdd : E)‖₊ = ‖x‖₊ := rfl @[simp] lemma nnnorm_ofAdd (x : E) : ‖ofAdd x‖₊ = ‖x‖₊ := rfl end NNNorm instance Additive.seminormedAddGroup [SeminormedGroup E] : SeminormedAddGroup (Additive E) where - dist_eq x y := dist_eq_norm_div (toMul x) (toMul y) + dist_eq x y := dist_eq_norm_div x.toMul y.toMul instance Multiplicative.seminormedGroup [SeminormedAddGroup E] : SeminormedGroup (Multiplicative E) where - dist_eq x y := dist_eq_norm_sub (toMul x) (toMul y) + dist_eq x y := dist_eq_norm_sub x.toAdd y.toAdd instance Additive.seminormedCommGroup [SeminormedCommGroup E] : SeminormedAddCommGroup (Additive E) := diff --git a/Mathlib/Data/Complex/Exponential.lean b/Mathlib/Data/Complex/Exponential.lean index 1054c7a046c4b..f789146aa2884 100644 --- a/Mathlib/Data/Complex/Exponential.lean +++ b/Mathlib/Data/Complex/Exponential.lean @@ -188,7 +188,7 @@ theorem exp_add : exp (x + y) = exp x * exp y := by /-- the exponential function as a monoid hom from `Multiplicative ℂ` to `ℂ` -/ @[simps] noncomputable def expMonoidHom : MonoidHom (Multiplicative ℂ) ℂ := - { toFun := fun z => exp (Multiplicative.toAdd z), + { toFun := fun z => exp z.toAdd, map_one' := by simp, map_mul' := by simp [exp_add] } @@ -689,7 +689,7 @@ nonrec theorem exp_add : exp (x + y) = exp x * exp y := by simp [exp_add, exp] /-- the exponential function as a monoid hom from `Multiplicative ℝ` to `ℝ` -/ @[simps] noncomputable def expMonoidHom : MonoidHom (Multiplicative ℝ) ℝ := - { toFun := fun x => exp (Multiplicative.toAdd x), + { toFun := fun x => exp x.toAdd, map_one' := by simp, map_mul' := by simp [exp_add] } diff --git a/Mathlib/Data/Int/Cast/Lemmas.lean b/Mathlib/Data/Int/Cast/Lemmas.lean index edf29c1a81e49..d887d52523d54 100644 --- a/Mathlib/Data/Int/Cast/Lemmas.lean +++ b/Mathlib/Data/Int/Cast/Lemmas.lean @@ -295,14 +295,14 @@ lemma zmultiplesHom_apply (x : β) (n : ℤ) : zmultiplesHom β x n = n • x := lemma zmultiplesHom_symm_apply (f : ℤ →+ β) : (zmultiplesHom β).symm f = f 1 := rfl @[to_additive existing (attr := simp)] -lemma zpowersHom_apply (x : α) (n : Multiplicative ℤ) : zpowersHom α x n = x ^ toAdd n := rfl +lemma zpowersHom_apply (x : α) (n : Multiplicative ℤ) : zpowersHom α x n = x ^ n.toAdd := rfl @[to_additive existing (attr := simp)] lemma zpowersHom_symm_apply (f : Multiplicative ℤ →* α) : (zpowersHom α).symm f = f (ofAdd 1) := rfl lemma MonoidHom.apply_mint (f : Multiplicative ℤ →* α) (n : Multiplicative ℤ) : - f n = f (ofAdd 1) ^ (toAdd n) := by + f n = f (ofAdd 1) ^ n.toAdd := by rw [← zpowersHom_symm_apply, ← zpowersHom_apply, Equiv.apply_symm_apply] lemma AddMonoidHom.apply_int (f : ℤ →+ β) (n : ℤ) : f n = n • f 1 := by @@ -324,7 +324,7 @@ def zpowersMulHom : α ≃* (Multiplicative ℤ →* α) := variable {α} @[simp] -lemma zpowersMulHom_apply (x : α) (n : Multiplicative ℤ) : zpowersMulHom α x n = x ^ toAdd n := rfl +lemma zpowersMulHom_apply (x : α) (n : Multiplicative ℤ) : zpowersMulHom α x n = x ^ n.toAdd := rfl @[simp] lemma zpowersMulHom_symm_apply (f : Multiplicative ℤ →* α) : diff --git a/Mathlib/Data/Int/WithZero.lean b/Mathlib/Data/Int/WithZero.lean index 5cf5aebd6da66..de4c3c4143538 100644 --- a/Mathlib/Data/Int/WithZero.lean +++ b/Mathlib/Data/Int/WithZero.lean @@ -14,7 +14,7 @@ the morphism `WithZeroMultInt.toNNReal`. ## Main Definitions * `WithZeroMultInt.toNNReal` : The `MonoidWithZeroHom` from `ℤₘ₀ → ℝ≥0` sending `0 ↦ 0` and - `x ↦ e^(Multiplicative.toAdd (WithZero.unzero hx)` when `x ≠ 0`, for a nonzero `e : ℝ≥0`. + `x ↦ e^((WithZero.unzero hx).toAdd)` when `x ≠ 0`, for a nonzero `e : ℝ≥0`. ## Main Results @@ -37,9 +37,9 @@ open Multiplicative WithZero namespace WithZeroMulInt /-- Given a nonzero `e : ℝ≥0`, this is the map `ℤₘ₀ → ℝ≥0` sending `0 ↦ 0` and - `x ↦ e^(Multiplicative.toAdd (WithZero.unzero hx)` when `x ≠ 0` as a `MonoidWithZeroHom`. -/ + `x ↦ e^(WithZero.unzero hx).toAdd` when `x ≠ 0` as a `MonoidWithZeroHom`. -/ def toNNReal {e : ℝ≥0} (he : e ≠ 0) : ℤₘ₀ →*₀ ℝ≥0 where - toFun := fun x ↦ if hx : x = 0 then 0 else e ^ Multiplicative.toAdd (WithZero.unzero hx) + toFun := fun x ↦ if hx : x = 0 then 0 else e ^ (WithZero.unzero hx).toAdd map_zero' := rfl map_one' := by simp only [dif_neg one_ne_zero] @@ -63,7 +63,7 @@ theorem toNNReal_pos_apply {e : ℝ≥0} (he : e ≠ 0) {x : ℤₘ₀} (hx : x split_ifs; rfl theorem toNNReal_neg_apply {e : ℝ≥0} (he : e ≠ 0) {x : ℤₘ₀} (hx : x ≠ 0) : - toNNReal he x = e ^ Multiplicative.toAdd (WithZero.unzero hx) := by + toNNReal he x = e ^ (WithZero.unzero hx).toAdd := by simp only [toNNReal, MonoidWithZeroHom.coe_mk, ZeroHom.coe_mk] split_ifs · tauto diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index d9fd72598d7bf..63025e62d2769 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -222,7 +222,7 @@ lemma multiplesHom_apply (x : β) (n : ℕ) : multiplesHom β x n = n • x := r @[to_additive existing (attr := simp)] lemma powersHom_apply (x : α) (n : Multiplicative ℕ) : - powersHom α x n = x ^ Multiplicative.toAdd n := rfl + powersHom α x n = x ^ n.toAdd := rfl lemma multiplesHom_symm_apply (f : ℕ →+ β) : (multiplesHom β).symm f = f 1 := rfl @@ -231,7 +231,7 @@ lemma powersHom_symm_apply (f : Multiplicative ℕ →* α) : (powersHom α).symm f = f (Multiplicative.ofAdd 1) := rfl lemma MonoidHom.apply_mnat (f : Multiplicative ℕ →* α) (n : Multiplicative ℕ) : - f n = f (Multiplicative.ofAdd 1) ^ (Multiplicative.toAdd n) := by + f n = f (Multiplicative.ofAdd 1) ^ n.toAdd := by rw [← powersHom_symm_apply, ← powersHom_apply, Equiv.apply_symm_apply] @[ext] @@ -258,7 +258,7 @@ def powersMulHom : α ≃* (Multiplicative ℕ →* α) := @[simp] lemma multiplesAddHom_apply (x : β) (n : ℕ) : multiplesAddHom β x n = n • x := rfl @[simp] -lemma powersMulHom_apply (x : α) (n : Multiplicative ℕ) : powersMulHom α x n = x ^ toAdd n := rfl +lemma powersMulHom_apply (x : α) (n : Multiplicative ℕ) : powersMulHom α x n = x ^ n.toAdd := rfl @[simp] lemma multiplesAddHom_symm_apply (f : ℕ →+ β) : (multiplesAddHom β).symm f = f 1 := rfl diff --git a/Mathlib/Data/ZMod/IntUnitsPower.lean b/Mathlib/Data/ZMod/IntUnitsPower.lean index 3927a7e59bb96..4b5cace9f38cd 100644 --- a/Mathlib/Data/ZMod/IntUnitsPower.lean +++ b/Mathlib/Data/ZMod/IntUnitsPower.lean @@ -24,7 +24,7 @@ by using `Module R (Additive M)` in its place, especially since this already has -/ instance : SMul (ZMod 2) (Additive ℤˣ) where - smul z au := .ofMul <| Additive.toMul au ^ z.val + smul z au := .ofMul <| au.toMul ^ z.val lemma ZMod.smul_units_def (z : ZMod 2) (au : Additive ℤˣ) : z • au = z.val • au := rfl @@ -34,7 +34,7 @@ lemma ZMod.natCast_smul_units (n : ℕ) (au : Additive ℤˣ) : (n : ZMod 2) • /-- This is an indirect way of saying that `ℤˣ` has a power operation by `ZMod 2`. -/ instance : Module (ZMod 2) (Additive ℤˣ) where - smul z au := .ofMul <| Additive.toMul au ^ z.val + smul z au := .ofMul <| au.toMul ^ z.val one_smul _ := Additive.toMul.injective <| pow_one _ mul_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul] @@ -44,7 +44,7 @@ instance : Module (ZMod 2) (Additive ℤˣ) where add_smul z₁ z₂ au := Additive.toMul.injective <| by dsimp only [ZMod.smul_units_def, toMul_nsmul, toMul_add] rw [← pow_add, ZMod.val_add, ← Int.units_pow_eq_pow_mod_two] - zero_smul au := Additive.toMul.injective <| pow_zero (Additive.toMul au) + zero_smul au := Additive.toMul.injective <| pow_zero au.toMul section CommSemiring variable {R : Type*} [CommSemiring R] [Module R (Additive ℤˣ)] @@ -55,7 +55,7 @@ In lemma names, this operations is called `uzpow` to match `zpow`. Notably this is satisfied by `R ∈ {ℕ, ℤ, ZMod 2}`. -/ instance Int.instUnitsPow : Pow ℤˣ R where - pow u r := Additive.toMul (r • Additive.ofMul u) + pow u r := (r • Additive.ofMul u).toMul -- The above instances form no typeclass diamonds with the standard power operators -- but we will need `reducible_and_instances` which currently fails https://github.com/leanprover-community/mathlib4/issues/10906 @@ -65,10 +65,10 @@ example : Int.instUnitsPow = DivInvMonoid.Pow := rfl @[simp] lemma ofMul_uzpow (u : ℤˣ) (r : R) : Additive.ofMul (u ^ r) = r • Additive.ofMul u := rfl @[simp] lemma toMul_uzpow (u : Additive ℤˣ) (r : R) : - Additive.toMul (r • u) = Additive.toMul u ^ r := rfl + (r • u).toMul = u.toMul ^ r := rfl @[norm_cast] lemma uzpow_natCast (u : ℤˣ) (n : ℕ) : u ^ (n : R) = u ^ n := by - change Additive.toMul ((n : R) • Additive.ofMul u) = _ + change ((n : R) • Additive.ofMul u).toMul = _ rw [Nat.cast_smul_eq_nsmul, toMul_nsmul, toMul_ofMul] -- See note [no_index around OfNat.ofNat] @@ -106,7 +106,7 @@ lemma uzpow_neg (s : ℤˣ) (x : R) : s ^ (-x) = (s ^ x)⁻¹ := Additive.ofMul.injective <| neg_smul x (Additive.ofMul s) @[norm_cast] lemma uzpow_intCast (u : ℤˣ) (z : ℤ) : u ^ (z : R) = u ^ z := by - change Additive.toMul ((z : R) • Additive.ofMul u) = _ + change ((z : R) • Additive.ofMul u).toMul = _ rw [Int.cast_smul_eq_zsmul, toMul_zsmul, toMul_ofMul] end CommRing diff --git a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean index 0b03b2f6c460d..e403b3b59936a 100644 --- a/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean +++ b/Mathlib/Dynamics/Circle/RotationNumber/TranslationNumber.lean @@ -253,7 +253,7 @@ theorem commute_iff_commute {f g : CircleDeg1Lift} : Commute f g ↔ Function.Co `translation (Multiplicative.ofAdd x)`. -/ def translate : Multiplicative ℝ →* CircleDeg1Liftˣ := MonoidHom.toHomUnits <| { toFun := fun x => - ⟨⟨fun y => Multiplicative.toAdd x + y, fun _ _ h => add_le_add_left h _⟩, fun _ => + ⟨⟨fun y => x.toAdd + y, fun _ _ h => add_le_add_left h _⟩, fun _ => (add_assoc _ _ _).symm⟩ map_one' := ext <| zero_add map_mul' := fun _ _ => ext <| add_assoc _ _ } diff --git a/Mathlib/GroupTheory/FiniteAbelian/Duality.lean b/Mathlib/GroupTheory/FiniteAbelian/Duality.lean index 308c47949bf39..dafdda97bea03 100644 --- a/Mathlib/GroupTheory/FiniteAbelian/Duality.lean +++ b/Mathlib/GroupTheory/FiniteAbelian/Duality.lean @@ -42,9 +42,9 @@ lemma exists_apply_ne_one_aux (H : ∀ n : ℕ, n ∣ Monoid.exponent G → ∀ obtain ⟨i, hi⟩ : ∃ i : ι, e a i ≠ 1 := by contrapose! ha exact (MulEquiv.map_eq_one_iff e).mp <| funext ha - have hi : Multiplicative.toAdd (e a i) ≠ 0 := by + have hi : (e a i).toAdd ≠ 0 := by simp only [ne_eq, toAdd_eq_zero, hi, not_false_eq_true] - obtain ⟨φi, hφi⟩ := H (n i) (dvd_exponent e i) (Multiplicative.toAdd <| e a i) hi + obtain ⟨φi, hφi⟩ := H (n i) (dvd_exponent e i) ((e a i).toAdd) hi use (φi.comp (Pi.evalMonoidHom (fun (i : ι) ↦ Multiplicative (ZMod (n i))) i)).comp e simpa only [coe_comp, coe_coe, Function.comp_apply, Pi.evalMonoidHom_apply, ne_eq] using hφi diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 20adcbf583239..63e23ebbb7c77 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -422,7 +422,7 @@ theorem constVAdd_symm (v : V₁) : (constVAdd k P₁ v).symm = constVAdd k P₁ /-- A more bundled version of `AffineEquiv.constVAdd`. -/ @[simps] def constVAddHom : Multiplicative V₁ →* P₁ ≃ᵃ[k] P₁ where - toFun v := constVAdd k P₁ (Multiplicative.toAdd v) + toFun v := constVAdd k P₁ v.toAdd map_one' := constVAdd_zero _ _ map_mul' := constVAdd_add _ P₁ diff --git a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean index 31c16252e0c0d..4fbfe52301fe9 100644 --- a/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean +++ b/Mathlib/NumberTheory/NumberField/Units/DirichletTheorem.lean @@ -75,7 +75,7 @@ variable (K) /-- The logarithmic embedding of the units (seen as an `Additive` group). -/ def _root_.NumberField.Units.logEmbedding : Additive ((𝓞 K)ˣ) →+ ({w : InfinitePlace K // w ≠ w₀} → ℝ) := -{ toFun := fun x w => mult w.val * Real.log (w.val ↑(Additive.toMul x)) +{ toFun := fun x w => mult w.val * Real.log (w.val ↑x.toMul) map_zero' := by simp; rfl map_add' := fun _ _ => by simp [Real.log_mul, mul_add]; rfl } @@ -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 ((basisModTorsion K i).toMul:) theorem fundSystem_mk (i : Fin (rank K)) : Additive.ofMul (QuotientGroup.mk (fundSystem K i)) = (basisModTorsion K i) := by diff --git a/Mathlib/RepresentationTheory/Basic.lean b/Mathlib/RepresentationTheory/Basic.lean index fbac878c2ff51..f17cfb213b7e5 100644 --- a/Mathlib/RepresentationTheory/Basic.lean +++ b/Mathlib/RepresentationTheory/Basic.lean @@ -292,7 +292,7 @@ def ofMulDistribMulAction : Representation ℤ M (Additive G) := ((monoidEndToAdditive G : _ →* _).comp (MulDistribMulAction.toMonoidEnd M G)) @[simp] theorem ofMulDistribMulAction_apply_apply (g : M) (a : Additive G) : - ofMulDistribMulAction M G g a = Additive.ofMul (g • Additive.toMul a) := rfl + ofMulDistribMulAction M G g a = Additive.ofMul (g • a.toMul) := rfl end MulDistribMulAction section Group diff --git a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean index b05641a81a752..24d8d8f5bd37c 100644 --- a/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean +++ b/Mathlib/RepresentationTheory/GroupCohomology/LowDegree.lean @@ -710,7 +710,7 @@ theorem H1LequivOfIsTrivial_comp_H1_π [A.IsTrivial] : @[simp] theorem H1LequivOfIsTrivial_H1_π_apply_apply [A.IsTrivial] (f : oneCocycles A) (x : Additive G) : - H1LequivOfIsTrivial A (H1_π A f) x = f (Additive.toMul x) := rfl + H1LequivOfIsTrivial A (H1_π A f) x = f x.toMul := rfl @[simp] theorem H1LequivOfIsTrivial_symm_apply [A.IsTrivial] (f : Additive G →+ A) : (H1LequivOfIsTrivial A).symm f = H1_π A ((oneCocyclesLequivOfIsTrivial A).symm f) := diff --git a/Mathlib/RepresentationTheory/Rep.lean b/Mathlib/RepresentationTheory/Rep.lean index d4ddfa12b4c67..0a14328b88a96 100644 --- a/Mathlib/RepresentationTheory/Rep.lean +++ b/Mathlib/RepresentationTheory/Rep.lean @@ -254,7 +254,7 @@ variable (M G : Type) [Monoid M] [CommGroup G] [MulDistribMulAction M G] def ofMulDistribMulAction : Rep ℤ M := Rep.of (Representation.ofMulDistribMulAction M G) @[simp] theorem ofMulDistribMulAction_ρ_apply_apply (g : M) (a : Additive G) : - (ofMulDistribMulAction M G).ρ g a = Additive.ofMul (g • Additive.toMul a) := rfl + (ofMulDistribMulAction M G).ρ g a = Additive.ofMul (g • a.toMul) := rfl /-- Given an `R`-algebra `S`, the `ℤ`-linear representation associated to the natural action of `S ≃ₐ[R] S` on `Sˣ`. -/ diff --git a/Mathlib/RingTheory/FreeCommRing.lean b/Mathlib/RingTheory/FreeCommRing.lean index 6efc6c11c1cd6..e1b0736aa1404 100644 --- a/Mathlib/RingTheory/FreeCommRing.lean +++ b/Mathlib/RingTheory/FreeCommRing.lean @@ -134,7 +134,7 @@ private def liftToMultiset : (α → R) ≃ (Multiplicative (Multiset α) →* R left_inv f := funext fun x => show (Multiset.map f {x}).prod = _ by simp right_inv F := MonoidHom.ext fun x => let F' := MonoidHom.toAdditive'' F - let x' := Multiplicative.toAdd x + let x' := x.toAdd show (Multiset.map (fun a => F' {a}) x').sum = F' x' by erw [← Multiset.map_map (fun x => F' x) (fun x => {x}), ← AddMonoidHom.map_multiset_sum] exact DFunLike.congr_arg F (Multiset.sum_map_singleton x') diff --git a/Mathlib/RingTheory/LaurentSeries.lean b/Mathlib/RingTheory/LaurentSeries.lean index f8ee85f00a0c0..1e09014fa6640 100644 --- a/Mathlib/RingTheory/LaurentSeries.lean +++ b/Mathlib/RingTheory/LaurentSeries.lean @@ -869,7 +869,7 @@ theorem Cauchy.eventually_mem_nhds {ℱ : Filter K⸨X⸩} (hℱ : Cauchy ℱ) obtain ⟨γ, hU₁⟩ := Valued.mem_nhds.mp hU suffices ∀ᶠ f in ℱ, f ∈ {y : K⸨X⸩ | Valued.v (y - limit hℱ) < ↑γ} by apply this.mono fun _ hf ↦ hU₁ hf - set D := -(Multiplicative.toAdd (WithZero.unzero γ.ne_zero) - 1) with hD₀ + set D := -((WithZero.unzero γ.ne_zero).toAdd - 1) with hD₀ have hD : ((Multiplicative.ofAdd (-D) : Multiplicative ℤ) : ℤₘ₀) < γ := by rw [← WithZero.coe_unzero γ.ne_zero, WithZero.coe_lt_coe, hD₀, neg_neg, ofAdd_sub, ofAdd_toAdd, div_lt_comm, div_self', ← ofAdd_zero, Multiplicative.ofAdd_lt] diff --git a/Mathlib/Topology/Constructions.lean b/Mathlib/Topology/Constructions.lean index 7d92b1e682ec5..280967dfdc29d 100644 --- a/Mathlib/Topology/Constructions.lean +++ b/Mathlib/Topology/Constructions.lean @@ -115,9 +115,9 @@ theorem nhds_ofMul (x : X) : 𝓝 (ofMul x) = map ofMul (𝓝 x) := rfl theorem nhds_ofAdd (x : X) : 𝓝 (ofAdd x) = map ofAdd (𝓝 x) := rfl -theorem nhds_toMul (x : Additive X) : 𝓝 (toMul x) = map toMul (𝓝 x) := rfl +theorem nhds_toMul (x : Additive X) : 𝓝 x.toMul = map toMul (𝓝 x) := rfl -theorem nhds_toAdd (x : Multiplicative X) : 𝓝 (toAdd x) = map toAdd (𝓝 x) := rfl +theorem nhds_toAdd (x : Multiplicative X) : 𝓝 x.toAdd = map toAdd (𝓝 x) := rfl end diff --git a/Mathlib/Topology/EMetricSpace/Defs.lean b/Mathlib/Topology/EMetricSpace/Defs.lean index 445f6803e42d1..b791062fbf650 100644 --- a/Mathlib/Topology/EMetricSpace/Defs.lean +++ b/Mathlib/Topology/EMetricSpace/Defs.lean @@ -657,11 +657,11 @@ theorem edist_ofAdd (a b : X) : edist (ofAdd a) (ofAdd b) = edist a b := rfl @[simp] -theorem edist_toMul (a b : Additive X) : edist (toMul a) (toMul b) = edist a b := +theorem edist_toMul (a b : Additive X) : edist a.toMul b.toMul = edist a b := rfl @[simp] -theorem edist_toAdd (a b : Multiplicative X) : edist (toAdd a) (toAdd b) = edist a b := +theorem edist_toAdd (a b : Multiplicative X) : edist a.toAdd b.toAdd = edist a b := rfl end diff --git a/Mathlib/Topology/MetricSpace/Defs.lean b/Mathlib/Topology/MetricSpace/Defs.lean index 93866859fadb0..e678bdcfedb6b 100644 --- a/Mathlib/Topology/MetricSpace/Defs.lean +++ b/Mathlib/Topology/MetricSpace/Defs.lean @@ -197,9 +197,9 @@ instance : Dist (Multiplicative X) := ‹Dist X› @[simp] theorem dist_ofAdd (a b : X) : dist (ofAdd a) (ofAdd b) = dist a b := rfl -@[simp] theorem dist_toMul (a b : Additive X) : dist (toMul a) (toMul b) = dist a b := rfl +@[simp] theorem dist_toMul (a b : Additive X) : dist a.toMul b.toMul = dist a b := rfl -@[simp] theorem dist_toAdd (a b : Multiplicative X) : dist (toAdd a) (toAdd b) = dist a b := rfl +@[simp] theorem dist_toAdd (a b : Multiplicative X) : dist a.toAdd b.toAdd = dist a b := rfl end @@ -211,10 +211,10 @@ variable [PseudoMetricSpace X] @[simp] theorem nndist_ofAdd (a b : X) : nndist (ofAdd a) (ofAdd b) = nndist a b := rfl -@[simp] theorem nndist_toMul (a b : Additive X) : nndist (toMul a) (toMul b) = nndist a b := rfl +@[simp] theorem nndist_toMul (a b : Additive X) : nndist a.toMul b.toMul = nndist a b := rfl @[simp] -theorem nndist_toAdd (a b : Multiplicative X) : nndist (toAdd a) (toAdd b) = nndist a b := rfl +theorem nndist_toAdd (a b : Multiplicative X) : nndist a.toAdd b.toAdd = nndist a b := rfl end diff --git a/Mathlib/Topology/MetricSpace/IsometricSMul.lean b/Mathlib/Topology/MetricSpace/IsometricSMul.lean index 8bbc25f34f0ab..7af5ce1593cca 100644 --- a/Mathlib/Topology/MetricSpace/IsometricSMul.lean +++ b/Mathlib/Topology/MetricSpace/IsometricSMul.lean @@ -420,26 +420,26 @@ instance Pi.isometricSMul'' {ι} {M : ι → Type*} [Fintype ι] [∀ i, Mul (M ⟨fun c => .piMap (fun i (x : M i) => x * c.unop i) fun _ => isometry_mul_right _⟩ instance Additive.isometricVAdd : IsometricVAdd (Additive M) X := - ⟨fun c => isometry_smul X (toMul c)⟩ + ⟨fun c => isometry_smul X c.toMul⟩ instance Additive.isometricVAdd' [Mul M] [PseudoEMetricSpace M] [IsometricSMul M M] : IsometricVAdd (Additive M) (Additive M) := - ⟨fun c x y => edist_smul_left (toMul c) (toMul x) (toMul y)⟩ + ⟨fun c x y => edist_smul_left c.toMul x.toMul y.toMul⟩ instance Additive.isometricVAdd'' [Mul M] [PseudoEMetricSpace M] [IsometricSMul Mᵐᵒᵖ M] : IsometricVAdd (Additive M)ᵃᵒᵖ (Additive M) := - ⟨fun c x y => edist_smul_left (MulOpposite.op (toMul c.unop)) (toMul x) (toMul y)⟩ + ⟨fun c x y => edist_smul_left (MulOpposite.op c.unop.toMul) x.toMul y.toMul⟩ instance Multiplicative.isometricSMul {M X} [VAdd M X] [PseudoEMetricSpace X] [IsometricVAdd M X] : IsometricSMul (Multiplicative M) X := - ⟨fun c => isometry_vadd X (toAdd c)⟩ + ⟨fun c => isometry_vadd X c.toAdd⟩ instance Multiplicative.isometricSMul' [Add M] [PseudoEMetricSpace M] [IsometricVAdd M M] : IsometricSMul (Multiplicative M) (Multiplicative M) := - ⟨fun c x y => edist_vadd_left (toAdd c) (toAdd x) (toAdd y)⟩ + ⟨fun c x y => edist_vadd_left c.toAdd x.toAdd y.toAdd⟩ instance Multiplicative.isometricVAdd'' [Add M] [PseudoEMetricSpace M] [IsometricVAdd Mᵃᵒᵖ M] : IsometricSMul (Multiplicative M)ᵐᵒᵖ (Multiplicative M) := - ⟨fun c x y => edist_vadd_left (AddOpposite.op (toAdd c.unop)) (toAdd x) (toAdd y)⟩ + ⟨fun c x y => edist_vadd_left (AddOpposite.op c.unop.toAdd) x.toAdd y.toAdd⟩ end Instances From c9cb88e7c1a95c987df4d3d68dd55f4c90d5bf11 Mon Sep 17 00:00:00 2001 From: grunweg Date: Sat, 23 Nov 2024 00:14:53 +0000 Subject: [PATCH 046/250] chore(workflows): add another missing job name (#19386) Analogous to #19382. --- .github/workflows/sync_closed_tasks.yaml | 1 + 1 file changed, 1 insertion(+) diff --git a/.github/workflows/sync_closed_tasks.yaml b/.github/workflows/sync_closed_tasks.yaml index 9c73e7e8a24c6..6c5383d62099c 100644 --- a/.github/workflows/sync_closed_tasks.yaml +++ b/.github/workflows/sync_closed_tasks.yaml @@ -12,6 +12,7 @@ on: jobs: build: + name: "Cross off linked issues" runs-on: ubuntu-latest steps: - name: Cross off any linked issue and PR references From 5483f1cf4b7d47e8843e2f4b71605bc74c60a3a7 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 23 Nov 2024 02:13:37 +0000 Subject: [PATCH 047/250] chore: update Mathlib dependencies 2024-11-23 (#19389) 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 9d7aff8b1542f..762bb1280e3ba 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "33d7f346440869364a2cd077bde8cebf73243aaa", + "rev": "8b52587ff32e2e443cce6109b5305341289339e7", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 5b58f54371168aba52c6d88e02af49b56ee41d6f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ga=C3=ABtan=20Serr=C3=A9?= Date: Sat, 23 Nov 2024 04:20:30 +0000 Subject: [PATCH 048/250] feat(Order/Filter/Basic): eventually_iff_all_subsets and specializations (#13275) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit The set `{x | p x}` belongs to `f` iff `∀ s, {x | x ∈ s → p x} ∈ f`. Specialization for `EventuallyEq` and `Eventually.LE`. Suggested by @urkud. Co-authored-by: Gaëtan Serré <56162277+gaetanserre@users.noreply.github.com> Co-authored-by: Johan Commelin --- Mathlib/Order/Filter/Basic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index ee7fd636ed470..995ddfea3234b 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -895,6 +895,11 @@ theorem eventually_inf_principal {f : Filter α} {p : α → Prop} {s : Set α} (∀ᶠ x in f ⊓ 𝓟 s, p x) ↔ ∀ᶠ x in f, x ∈ s → p x := mem_inf_principal +theorem eventually_iff_all_subsets {f : Filter α} {p : α → Prop} : + (∀ᶠ x in f, p x) ↔ ∀ (s : Set α), ∀ᶠ x in f, x ∈ s → p x where + mp h _ := by filter_upwards [h] with _ pa _ using pa + mpr h := by filter_upwards [h univ] with _ pa using pa (by simp) + /-! ### Frequently -/ theorem Eventually.frequently {f : Filter α} [NeBot f] {p : α → Prop} (h : ∀ᶠ x in f, p x) : @@ -1224,6 +1229,10 @@ theorem eventuallyEq_iff_sub [AddGroup β] {f g : α → β} {l : Filter α} : f =ᶠ[l] g ↔ f - g =ᶠ[l] 0 := ⟨fun h => h.sub_eq, fun h => by simpa using h.add (EventuallyEq.refl l g)⟩ +theorem eventuallyEq_iff_all_subsets {f g : α → β} {l : Filter α} : + f =ᶠ[l] g ↔ ∀ s : Set α, ∀ᶠ x in l, x ∈ s → f x = g x := + eventually_iff_all_subsets + section LE variable [LE β] {l : Filter α} @@ -1236,6 +1245,10 @@ theorem eventuallyLE_congr {f f' g g' : α → β} (hf : f =ᶠ[l] f') (hg : g = f ≤ᶠ[l] g ↔ f' ≤ᶠ[l] g' := ⟨fun H => H.congr hf hg, fun H => H.congr hf.symm hg.symm⟩ +theorem eventuallyLE_iff_all_subsets {f g : α → β} {l : Filter α} : + f ≤ᶠ[l] g ↔ ∀ s : Set α, ∀ᶠ x in l, x ∈ s → f x ≤ g x := + eventually_iff_all_subsets + end LE section Preorder From cff32a812b0c1e24ee480674e97160d1070b7393 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 04:49:22 +0000 Subject: [PATCH 049/250] chore: adaptations for nightly-2024-11-22 (#19390) Co-authored-by: Kyle Miller Co-authored-by: JovanGerb Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Johan Commelin --- Mathlib/Data/List/Indexes.lean | 1 - Mathlib/SetTheory/Cardinal/Basic.lean | 1 - .../Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean | 2 +- lake-manifest.json | 4 ++-- lean-toolchain | 2 +- 5 files changed, 4 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 6c8ce5ff6c794..2097c5adf2cb5 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Jannis Limperg. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jannis Limperg -/ -import Mathlib.Order.Lattice import Mathlib.Data.List.Basic /-! diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index b309bcdad858b..2b76c69c48ddc 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -14,7 +14,6 @@ import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.InitialSeg import Mathlib.Order.SuccPred.CompleteLinearOrder import Mathlib.SetTheory.Cardinal.SchroederBernstein -import Mathlib.Algebra.Order.Group.Nat import Mathlib.Algebra.Order.GroupWithZero.Canonical import Mathlib.Algebra.Order.Ring.Canonical diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean index 9858e94c23f14..17812e9c56178 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean @@ -60,7 +60,7 @@ nonnegative. def checkSuccess : SimplexAlgorithmM matType Bool := do let lastIdx := (← get).free.size - 1 return (← get).mat[(0, lastIdx)]! > 0 && - (← Nat.allM (← get).basic.size (fun i => do return (← get).mat[(i, lastIdx)]! >= 0)) + (← (← get).basic.size.allM (fun i _ => do return (← get).mat[(i, lastIdx)]! ≥ 0)) /-- Chooses an entering variable: among the variables with a positive coefficient in the objective diff --git a/lake-manifest.json b/lake-manifest.json index 4820e7097b0d2..ec9d70d4be933 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "b3935d53ce8dad5665af7ac41f06f0f6de4d942f", + "rev": "4efc0bf8d39cd7fb92e59f44ad38091ca14de2ce", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "lean-pr-testing-6128", + "inputRev": "nightly-testing", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lean-toolchain b/lean-toolchain index f589f3bf8fa9a..ae2a1d4650fc0 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-21 +leanprover/lean4:nightly-2024-11-22 From e360f005e729a625b4f36183a23a57e1865a527d Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Sat, 23 Nov 2024 04:58:45 +0000 Subject: [PATCH 050/250] feat: add `Subgroup.index_antitone` and `Subgroup.index_strictAnti` (#19188) from flt-regular. --- Mathlib/GroupTheory/Index.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 8357feb56afc9..c31a41f62822f 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -554,6 +554,17 @@ variable {H K} theorem finiteIndex_of_le [FiniteIndex H] (h : H ≤ K) : FiniteIndex K := ⟨ne_zero_of_dvd_ne_zero FiniteIndex.finiteIndex (index_dvd_of_le h)⟩ +@[to_additive (attr := gcongr)] +lemma index_antitone (h : H ≤ K) [H.FiniteIndex] : K.index ≤ H.index := + Nat.le_of_dvd (Nat.zero_lt_of_ne_zero FiniteIndex.finiteIndex) (index_dvd_of_le h) + +@[to_additive (attr := gcongr)] +lemma index_strictAnti (h : H < K) [H.FiniteIndex] : K.index < H.index := by + have h0 : K.index ≠ 0 := (finiteIndex_of_le h.le).finiteIndex + apply lt_of_le_of_ne (index_antitone h.le) + rw [← relindex_mul_index h.le, Ne, eq_comm, mul_eq_right₀ h0, relindex_eq_one] + exact h.not_le + variable (H K) @[to_additive] From 506763e8dc880a44e92ebe195c4d369d3dcef5f0 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Sat, 23 Nov 2024 04:58:46 +0000 Subject: [PATCH 051/250] chore: rename *tower_top_of_injective lemmas (#19246) --- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 2 +- Mathlib/RingTheory/Algebraic.lean | 38 +++++++++++-------- Mathlib/RingTheory/AlgebraicIndependent.lean | 3 +- Mathlib/RingTheory/Localization/Integral.lean | 2 +- 4 files changed, 25 insertions(+), 20 deletions(-) diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index 2ec4275183cac..c2bc06cd1f067 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -410,7 +410,7 @@ noncomputable def equivOfEquivAux (hSR : S ≃+* R) : haveI : IsScalarTower S R L := IsScalarTower.of_algebraMap_eq (by simp [RingHom.algebraMap_toAlgebra]) haveI : NoZeroSMulDivisors R S := NoZeroSMulDivisors.of_algebraMap_injective hSR.symm.injective - have : Algebra.IsAlgebraic R L := (IsAlgClosure.isAlgebraic.tower_top_of_injective + have : Algebra.IsAlgebraic R L := (IsAlgClosure.isAlgebraic.extendScalars (show Function.Injective (algebraMap S R) from hSR.injective)) refine ⟨equivOfAlgebraic' R S L M, ?_⟩ diff --git a/Mathlib/RingTheory/Algebraic.lean b/Mathlib/RingTheory/Algebraic.lean index e8fe139139573..4752101e3862d 100644 --- a/Mathlib/RingTheory/Algebraic.lean +++ b/Mathlib/RingTheory/Algebraic.lean @@ -499,31 +499,34 @@ variable [Algebra R S] [Algebra S A] [Algebra R A] [IsScalarTower R S A] /-- If `x` is algebraic over `R`, then `x` is algebraic over `S` when `S` is an extension of `R`, and the map from `R` to `S` is injective. -/ -theorem IsAlgebraic.tower_top_of_injective - (hinj : Function.Injective (algebraMap R S)) {x : A} +theorem IsAlgebraic.extendScalars (hinj : Function.Injective (algebraMap R S)) {x : A} (A_alg : IsAlgebraic R x) : IsAlgebraic S x := let ⟨p, hp₁, hp₂⟩ := A_alg ⟨p.map (algebraMap _ _), by rwa [Ne, ← degree_eq_bot, degree_map_eq_of_injective hinj, degree_eq_bot], by simpa⟩ -/-- A special case of `IsAlgebraic.tower_top_of_injective`. This is extracted as a theorem - because in some cases `IsAlgebraic.tower_top_of_injective` will just runs out of memory. -/ +@[deprecated (since := "2024-11-18")] +alias IsAlgebraic.tower_top_of_injective := IsAlgebraic.extendScalars + +/-- A special case of `IsAlgebraic.extendScalars`. This is extracted as a theorem + because in some cases `IsAlgebraic.extendScalars` will just runs out of memory. -/ theorem IsAlgebraic.tower_top_of_subalgebra_le {A B : Subalgebra R S} (hle : A ≤ B) {x : S} (h : IsAlgebraic A x) : IsAlgebraic B x := by letI : Algebra A B := (Subalgebra.inclusion hle).toAlgebra haveI : IsScalarTower A B S := .of_algebraMap_eq fun _ ↦ rfl - exact h.tower_top_of_injective (Subalgebra.inclusion_injective hle) + exact h.extendScalars (Subalgebra.inclusion_injective hle) /-- If `x` is transcendental over `S`, then `x` is transcendental over `R` when `S` is an extension of `R`, and the map from `R` to `S` is injective. -/ -theorem Transcendental.of_tower_top_of_injective - (hinj : Function.Injective (algebraMap R S)) {x : A} - (h : Transcendental S x) : Transcendental R x := - fun H ↦ h (H.tower_top_of_injective hinj) +theorem Transcendental.restrictScalars (hinj : Function.Injective (algebraMap R S)) {x : A} + (h : Transcendental S x) : Transcendental R x := fun H ↦ h (H.extendScalars hinj) + +@[deprecated (since := "2024-11-18")] +alias Transcendental.of_tower_top_of_injective := Transcendental.restrictScalars -/-- A special case of `Transcendental.of_tower_top_of_injective`. This is extracted as a theorem - because in some cases `Transcendental.of_tower_top_of_injective` will just runs out of memory. -/ +/-- A special case of `Transcendental.restrictScalars`. This is extracted as a theorem + because in some cases `Transcendental.restrictScalars` will just runs out of memory. -/ theorem Transcendental.of_tower_top_of_subalgebra_le {A B : Subalgebra R S} (hle : A ≤ B) {x : S} (h : Transcendental B x) : Transcendental A x := @@ -531,9 +534,12 @@ theorem Transcendental.of_tower_top_of_subalgebra_le /-- If A is an algebraic algebra over R, then A is algebraic over S when S is an extension of R, and the map from `R` to `S` is injective. -/ -theorem Algebra.IsAlgebraic.tower_top_of_injective (hinj : Function.Injective (algebraMap R S)) +theorem Algebra.IsAlgebraic.extendScalars (hinj : Function.Injective (algebraMap R S)) [Algebra.IsAlgebraic R A] : Algebra.IsAlgebraic S A := - ⟨fun _ ↦ _root_.IsAlgebraic.tower_top_of_injective hinj (Algebra.IsAlgebraic.isAlgebraic _)⟩ + ⟨fun _ ↦ _root_.IsAlgebraic.extendScalars hinj (Algebra.IsAlgebraic.isAlgebraic _)⟩ + +@[deprecated (since := "2024-11-18")] +alias Algebra.IsAlgebraic.tower_top_of_injective := Algebra.IsAlgebraic.extendScalars theorem Algebra.IsAlgebraic.tower_bot_of_injective [Algebra.IsAlgebraic R A] (hinj : Function.Injective (algebraMap S A)) : @@ -553,7 +559,7 @@ variable (L) @[stacks 09GF "part one"] theorem IsAlgebraic.tower_top {x : A} (A_alg : IsAlgebraic K x) : IsAlgebraic L x := - A_alg.tower_top_of_injective (algebraMap K L).injective + A_alg.extendScalars (algebraMap K L).injective variable {L} (K) in /-- If `x` is transcendental over `L`, then `x` is transcendental over `K` when @@ -564,7 +570,7 @@ theorem Transcendental.of_tower_top {x : A} (h : Transcendental L x) : /-- If A is an algebraic algebra over K, then A is algebraic over L when L is an extension of K -/ @[stacks 09GF "part two"] theorem Algebra.IsAlgebraic.tower_top [Algebra.IsAlgebraic K A] : Algebra.IsAlgebraic L A := - Algebra.IsAlgebraic.tower_top_of_injective (algebraMap K L).injective + Algebra.IsAlgebraic.extendScalars (algebraMap K L).injective variable (K) @@ -928,7 +934,7 @@ theorem transcendental_supported_polynomial_aeval_X_iff exact Algebra.adjoin_mono (Set.singleton_subset_iff.2 (Set.mem_image_of_mem _ h)) (Polynomial.aeval_mem_adjoin_singleton _ _) · rw [← transcendental_polynomial_aeval_X_iff R i] - refine h.of_tower_top_of_injective fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ + refine h.restrictScalars fun _ _ heq ↦ MvPolynomial.C_injective σ R ?_ simp_rw [← MvPolynomial.algebraMap_eq] exact congr($(heq).1) diff --git a/Mathlib/RingTheory/AlgebraicIndependent.lean b/Mathlib/RingTheory/AlgebraicIndependent.lean index 6865dee085405..7109a5bc4b402 100644 --- a/Mathlib/RingTheory/AlgebraicIndependent.lean +++ b/Mathlib/RingTheory/AlgebraicIndependent.lean @@ -724,8 +724,7 @@ theorem IsTranscendenceBasis.isAlgebraic_field {F E : Type*} {x : ι → E} (Subalgebra.inclusion (IntermediateField.algebra_adjoin_le_adjoin F S)).toRingHom.toAlgebra haveI : IsScalarTower (adjoin F S) (IntermediateField.adjoin F S) E := IsScalarTower.of_algebraMap_eq (congrFun rfl) - exact Algebra.IsAlgebraic.tower_top_of_injective (R := adjoin F S) - (Subalgebra.inclusion_injective _) + exact Algebra.IsAlgebraic.extendScalars (R := adjoin F S) (Subalgebra.inclusion_injective _) section Field diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index d2b5cef36d5e2..07c55c2708e4b 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -366,7 +366,7 @@ theorem isAlgebraic_iff' [Field K] [IsDomain R] [IsDomain S] [Algebra R K] [Alge rw [div_eq_mul_inv] refine IsIntegral.mul ?_ ?_ · rw [← isAlgebraic_iff_isIntegral] - refine .tower_top_of_injective + refine .extendScalars (NoZeroSMulDivisors.algebraMap_injective R (FractionRing R)) ?_ exact .algebraMap (h a) · rw [← isAlgebraic_iff_isIntegral] From 8d7c6defa0cb5ee614f419923bcc14dbc2f66583 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Sat, 23 Nov 2024 04:58:47 +0000 Subject: [PATCH 052/250] feat: add `(Algebra|IntermediateField).lift_cardinalMk_adjoin_le` (#19327) Also add `FreeAlgebra.cardinalMk_eq_max_lift`, etc. parallel to `MvPolynomial`. Also fix the name `cardinal_lift_mk_le_max` -> `cardinalMk_le_max_lift` for `MvPolynomial`. --- Mathlib.lean | 1 + Mathlib/Algebra/FreeAlgebra/Cardinality.lean | 78 ++++++++++++++++++++ Mathlib/Algebra/MvPolynomial/Cardinal.lean | 11 ++- Mathlib/FieldTheory/Adjoin.lean | 26 ++++++- 4 files changed, 113 insertions(+), 3 deletions(-) create mode 100644 Mathlib/Algebra/FreeAlgebra/Cardinality.lean diff --git a/Mathlib.lean b/Mathlib.lean index fc3aa1d9991a6..d142e08e33749 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -223,6 +223,7 @@ import Mathlib.Algebra.Field.Subfield.Defs import Mathlib.Algebra.Field.ULift import Mathlib.Algebra.Free import Mathlib.Algebra.FreeAlgebra +import Mathlib.Algebra.FreeAlgebra.Cardinality import Mathlib.Algebra.FreeMonoid.Basic import Mathlib.Algebra.FreeMonoid.Count import Mathlib.Algebra.FreeMonoid.Symbols diff --git a/Mathlib/Algebra/FreeAlgebra/Cardinality.lean b/Mathlib/Algebra/FreeAlgebra/Cardinality.lean new file mode 100644 index 0000000000000..801dea8a9298c --- /dev/null +++ b/Mathlib/Algebra/FreeAlgebra/Cardinality.lean @@ -0,0 +1,78 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.Algebra.FreeAlgebra +import Mathlib.SetTheory.Cardinal.Free + +/-! +# Cardinality of free algebras + +This file contains some results about the cardinality of `FreeAlgebra`, +parallel to that of `MvPolynomial`. +-/ + +universe u v + +variable (R : Type u) [CommSemiring R] + +open Cardinal + +namespace FreeAlgebra + +variable (X : Type v) + +@[simp] +theorem cardinalMk_eq_max_lift [Nonempty X] [Nontrivial R] : + #(FreeAlgebra R X) = Cardinal.lift.{v} #R ⊔ Cardinal.lift.{u} #X ⊔ ℵ₀ := by + have hX := mk_freeMonoid X + haveI : Infinite (FreeMonoid X) := infinite_iff.2 (by simp [hX]) + rw [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra, + mk_finsupp_lift_of_infinite, hX, lift_max, lift_aleph0, sup_comm, ← sup_assoc] + +@[simp] +theorem cardinalMk_eq_lift [IsEmpty X] : #(FreeAlgebra R X) = Cardinal.lift.{v} #R := by + have := lift_mk_eq'.2 ⟨show (FreeMonoid X →₀ R) ≃ R from Equiv.finsuppUnique⟩ + rw [lift_id'.{u, v}, lift_umax] at this + rwa [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra] + +@[nontriviality] +theorem cardinalMk_eq_one [Subsingleton R] : #(FreeAlgebra R X) = 1 := by + rw [equivMonoidAlgebraFreeMonoid.toEquiv.cardinal_eq, MonoidAlgebra, mk_eq_one] + +theorem cardinalMk_le_max_lift : + #(FreeAlgebra R X) ≤ Cardinal.lift.{v} #R ⊔ Cardinal.lift.{u} #X ⊔ ℵ₀ := by + cases subsingleton_or_nontrivial R + · exact (cardinalMk_eq_one R X).trans_le (le_max_of_le_right one_le_aleph0) + cases isEmpty_or_nonempty X + · exact (cardinalMk_eq_lift R X).trans_le (le_max_of_le_left <| le_max_left _ _) + · exact (cardinalMk_eq_max_lift R X).le + +variable (X : Type u) + +theorem cardinalMk_eq_max [Nonempty X] [Nontrivial R] : #(FreeAlgebra R X) = #R ⊔ #X ⊔ ℵ₀ := by + simp + +theorem cardinalMk_eq [IsEmpty X] : #(FreeAlgebra R X) = #R := by + simp + +theorem cardinalMk_le_max : #(FreeAlgebra R X) ≤ #R ⊔ #X ⊔ ℵ₀ := by + simpa using cardinalMk_le_max_lift R X + +end FreeAlgebra + +namespace Algebra + +theorem lift_cardinalMk_adjoin_le {A : Type v} [Semiring A] [Algebra R A] (s : Set A) : + lift.{u} #(adjoin R s) ≤ lift.{v} #R ⊔ lift.{u} #s ⊔ ℵ₀ := by + have H := mk_range_le_lift (f := FreeAlgebra.lift R ((↑) : s → A)) + rw [lift_umax, lift_id'.{v, u}] at H + rw [Algebra.adjoin_eq_range_freeAlgebra_lift] + exact H.trans (FreeAlgebra.cardinalMk_le_max_lift R s) + +theorem cardinalMk_adjoin_le {A : Type u} [Semiring A] [Algebra R A] (s : Set A) : + #(adjoin R s) ≤ #R ⊔ #s ⊔ ℵ₀ := by + simpa using lift_cardinalMk_adjoin_le R s + +end Algebra diff --git a/Mathlib/Algebra/MvPolynomial/Cardinal.lean b/Mathlib/Algebra/MvPolynomial/Cardinal.lean index 3ee977de57b85..09b4c9e178b8c 100644 --- a/Mathlib/Algebra/MvPolynomial/Cardinal.lean +++ b/Mathlib/Algebra/MvPolynomial/Cardinal.lean @@ -40,7 +40,10 @@ theorem cardinalMk_eq_lift [IsEmpty σ] : #(MvPolynomial σ R) = Cardinal.lift.{ @[deprecated (since := "2024-11-10")] alias cardinal_mk_eq_lift := cardinalMk_eq_lift -theorem cardinal_lift_mk_le_max {σ : Type u} {R : Type v} [CommSemiring R] : #(MvPolynomial σ R) ≤ +@[nontriviality] +theorem cardinalMk_eq_one [Subsingleton R] : #(MvPolynomial σ R) = 1 := mk_eq_one _ + +theorem cardinalMk_le_max_lift {σ : Type u} {R : Type v} [CommSemiring R] : #(MvPolynomial σ R) ≤ max (max (Cardinal.lift.{u} #R) <| Cardinal.lift.{v} #σ) ℵ₀ := by cases subsingleton_or_nontrivial R · exact (mk_eq_one _).trans_le (le_max_of_le_right one_le_aleph0) @@ -48,6 +51,8 @@ theorem cardinal_lift_mk_le_max {σ : Type u} {R : Type v} [CommSemiring R] : #( · exact cardinalMk_eq_lift.trans_le (le_max_of_le_left <| le_max_left _ _) · exact cardinalMk_eq_max_lift.le +@[deprecated (since := "2024-11-21")] alias cardinal_lift_mk_le_max := cardinalMk_le_max_lift + end TwoUniverses variable {σ R : Type u} [CommSemiring R] @@ -57,10 +62,12 @@ theorem cardinalMk_eq_max [Nonempty σ] [Nontrivial R] : @[deprecated (since := "2024-11-10")] alias cardinal_mk_eq_max := cardinalMk_eq_max +theorem cardinalMk_eq [IsEmpty σ] : #(MvPolynomial σ R) = #R := by simp + /-- The cardinality of the multivariate polynomial ring, `MvPolynomial σ R` is at most the maximum of `#R`, `#σ` and `ℵ₀` -/ theorem cardinalMk_le_max : #(MvPolynomial σ R) ≤ max (max #R #σ) ℵ₀ := - cardinal_lift_mk_le_max.trans <| by rw [lift_id, lift_id] + cardinalMk_le_max_lift.trans <| by rw [lift_id, lift_id] @[deprecated (since := "2024-11-10")] alias cardinal_mk_le_max := cardinalMk_le_max diff --git a/Mathlib/FieldTheory/Adjoin.lean b/Mathlib/FieldTheory/Adjoin.lean index 9ccf8c539d1f9..51d8f9c940bb7 100644 --- a/Mathlib/FieldTheory/Adjoin.lean +++ b/Mathlib/FieldTheory/Adjoin.lean @@ -11,6 +11,7 @@ import Mathlib.LinearAlgebra.Dimension.FreeAndStrongRankCondition import Mathlib.RingTheory.Adjoin.Dimension import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.TensorProduct.Basic +import Mathlib.SetTheory.Cardinal.Subfield /-! # Adjoining Elements to Fields @@ -1680,4 +1681,27 @@ theorem liftAlgHom_fieldRange_eq_of_range_eq (hg : Function.Injective g) end IsFractionRing -set_option linter.style.longFile 1700 +namespace IntermediateField + +universe u v + +open Cardinal + +variable (F : Type u) [Field F] + +theorem lift_cardinalMk_adjoin_le {E : Type v} [Field E] [Algebra F E] (s : Set E) : + Cardinal.lift.{u} #(adjoin F s) ≤ Cardinal.lift.{v} #F ⊔ Cardinal.lift.{u} #s ⊔ ℵ₀ := by + rw [show ↥(adjoin F s) = (adjoin F s).toSubfield from rfl, adjoin_toSubfield] + apply (Cardinal.lift_le.mpr (Subfield.cardinalMk_closure_le_max _)).trans + rw [lift_max, sup_le_iff, lift_aleph0] + refine ⟨(Cardinal.lift_le.mpr ((mk_union_le _ _).trans <| add_le_max _ _)).trans ?_, le_sup_right⟩ + simp_rw [lift_max, lift_aleph0, sup_assoc] + exact sup_le_sup_right mk_range_le_lift _ + +theorem cardinalMk_adjoin_le {E : Type u} [Field E] [Algebra F E] (s : Set E) : + #(adjoin F s) ≤ #F ⊔ #s ⊔ ℵ₀ := by + simpa using lift_cardinalMk_adjoin_le F s + +end IntermediateField + +set_option linter.style.longFile 1800 From a7f12adce3dd754dccc68fb2c83ecb76842f81f0 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 05:51:58 +0000 Subject: [PATCH 053/250] chore(Order/FixedPoints): remove adaptation notes for OrderHom regarding lean4 issue #1910 (#19380) this PR removes 3 porting notes regarding lean4 issue leanprover/lean4#1910, and makes adaptations to mathlib to make use of the newly working notation. Concretely, that means that `OrderHom.lfp f` now is written `f.lfp`, and similar for `OrderHom.dual` and `OrderHom.gfp`. there might still be some `dual x` application left inside code with `open OrderHom`, but those are hard to filter for with grep, as there are many declarations named `dual` in different namespaces. --- Mathlib/Order/FixedPoints.lean | 98 +++++++++---------- Mathlib/Order/Hom/Basic.lean | 7 +- Mathlib/Order/Hom/Bounded.lean | 6 +- Mathlib/Order/Interval/Basic.lean | 4 +- .../Cardinal/SchroederBernstein.lean | 3 +- .../Ordinal/FixedPointApproximants.lean | 32 +++--- 6 files changed, 71 insertions(+), 79 deletions(-) diff --git a/Mathlib/Order/FixedPoints.lean b/Mathlib/Order/FixedPoints.lean index e6f40735b5ddc..20419434bbcf6 100644 --- a/Mathlib/Order/FixedPoints.lean +++ b/Mathlib/Order/FixedPoints.lean @@ -51,76 +51,73 @@ def gfp : (α →o α) →o α where toFun f := sSup { a | a ≤ f a } monotone' _ _ hle := sSup_le_sSup fun a ha => le_trans ha (hle a) -theorem lfp_le {a : α} (h : f a ≤ a) : lfp f ≤ a := +theorem lfp_le {a : α} (h : f a ≤ a) : f.lfp ≤ a := sInf_le h -theorem lfp_le_fixed {a : α} (h : f a = a) : lfp f ≤ a := +theorem lfp_le_fixed {a : α} (h : f a = a) : f.lfp ≤ a := f.lfp_le h.le -theorem le_lfp {a : α} (h : ∀ b, f b ≤ b → a ≤ b) : a ≤ lfp f := +theorem le_lfp {a : α} (h : ∀ b, f b ≤ b → a ≤ b) : a ≤ f.lfp := le_sInf h --- Porting note: for the rest of the file, replace the dot notation `_.lfp` with `lfp _` --- same for `_.gfp`, `_.dual` --- Probably related to https://github.com/leanprover/lean4/issues/1910 -theorem map_le_lfp {a : α} (ha : a ≤ lfp f) : f a ≤ lfp f := +theorem map_le_lfp {a : α} (ha : a ≤ f.lfp) : f a ≤ f.lfp := f.le_lfp fun _ hb => (f.mono <| le_sInf_iff.1 ha _ hb).trans hb @[simp] -theorem map_lfp : f (lfp f) = lfp f := - have h : f (lfp f) ≤ lfp f := f.map_le_lfp le_rfl +theorem map_lfp : f f.lfp = f.lfp := + have h : f f.lfp ≤ f.lfp := f.map_le_lfp le_rfl h.antisymm <| f.lfp_le <| f.mono h -theorem isFixedPt_lfp : IsFixedPt f (lfp f) := +theorem isFixedPt_lfp : IsFixedPt f f.lfp := f.map_lfp -theorem lfp_le_map {a : α} (ha : lfp f ≤ a) : lfp f ≤ f a := +theorem lfp_le_map {a : α} (ha : f.lfp ≤ a) : f.lfp ≤ f a := calc - lfp f = f (lfp f) := f.map_lfp.symm + f.lfp = f f.lfp := f.map_lfp.symm _ ≤ f a := f.mono ha -theorem isLeast_lfp_le : IsLeast { a | f a ≤ a } (lfp f) := +theorem isLeast_lfp_le : IsLeast { a | f a ≤ a } f.lfp := ⟨f.map_lfp.le, fun _ => f.lfp_le⟩ -theorem isLeast_lfp : IsLeast (fixedPoints f) (lfp f) := +theorem isLeast_lfp : IsLeast (fixedPoints f) f.lfp := ⟨f.isFixedPt_lfp, fun _ => f.lfp_le_fixed⟩ -theorem lfp_induction {p : α → Prop} (step : ∀ a, p a → a ≤ lfp f → p (f a)) - (hSup : ∀ s, (∀ a ∈ s, p a) → p (sSup s)) : p (lfp f) := by - set s := { a | a ≤ lfp f ∧ p a } +theorem lfp_induction {p : α → Prop} (step : ∀ a, p a → a ≤ f.lfp → p (f a)) + (hSup : ∀ s, (∀ a ∈ s, p a) → p (sSup s)) : p f.lfp := by + set s := { a | a ≤ f.lfp ∧ p a } specialize hSup s fun a => And.right - suffices sSup s = lfp f from this ▸ hSup - have h : sSup s ≤ lfp f := sSup_le fun b => And.left + suffices sSup s = f.lfp from this ▸ hSup + have h : sSup s ≤ f.lfp := sSup_le fun b => And.left have hmem : f (sSup s) ∈ s := ⟨f.map_le_lfp h, step _ hSup h⟩ exact h.antisymm (f.lfp_le <| le_sSup hmem) -theorem le_gfp {a : α} (h : a ≤ f a) : a ≤ gfp f := +theorem le_gfp {a : α} (h : a ≤ f a) : a ≤ f.gfp := le_sSup h -theorem gfp_le {a : α} (h : ∀ b, b ≤ f b → b ≤ a) : gfp f ≤ a := +theorem gfp_le {a : α} (h : ∀ b, b ≤ f b → b ≤ a) : f.gfp ≤ a := sSup_le h -theorem isFixedPt_gfp : IsFixedPt f (gfp f) := +theorem isFixedPt_gfp : IsFixedPt f f.gfp := f.dual.isFixedPt_lfp @[simp] -theorem map_gfp : f (gfp f) = gfp f := +theorem map_gfp : f f.gfp = f.gfp := f.dual.map_lfp -theorem map_le_gfp {a : α} (ha : a ≤ gfp f) : f a ≤ gfp f := +theorem map_le_gfp {a : α} (ha : a ≤ f.gfp) : f a ≤ f.gfp := f.dual.lfp_le_map ha -theorem gfp_le_map {a : α} (ha : gfp f ≤ a) : gfp f ≤ f a := +theorem gfp_le_map {a : α} (ha : f.gfp ≤ a) : f.gfp ≤ f a := f.dual.map_le_lfp ha -theorem isGreatest_gfp_le : IsGreatest { a | a ≤ f a } (gfp f) := +theorem isGreatest_gfp_le : IsGreatest { a | a ≤ f a } f.gfp := f.dual.isLeast_lfp_le -theorem isGreatest_gfp : IsGreatest (fixedPoints f) (gfp f) := +theorem isGreatest_gfp : IsGreatest (fixedPoints f) f.gfp := f.dual.isLeast_lfp -theorem gfp_induction {p : α → Prop} (step : ∀ a, p a → gfp f ≤ a → p (f a)) - (hInf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : p (gfp f) := +theorem gfp_induction {p : α → Prop} (step : ∀ a, p a → f.gfp ≤ a → p (f a)) + (hInf : ∀ s, (∀ a ∈ s, p a) → p (sInf s)) : p f.gfp := f.dual.lfp_induction step hInf end Basic @@ -130,27 +127,26 @@ section Eqn variable [CompleteLattice α] [CompleteLattice β] (f : β →o α) (g : α →o β) -- Rolling rule -theorem map_lfp_comp : f (lfp (g.comp f)) = lfp (f.comp g) := +theorem map_lfp_comp : f (g.comp f).lfp = (f.comp g).lfp := le_antisymm ((f.comp g).map_lfp ▸ f.mono (lfp_le_fixed _ <| congr_arg g (f.comp g).map_lfp)) <| lfp_le _ (congr_arg f (g.comp f).map_lfp).le -theorem map_gfp_comp : f (gfp (g.comp f)) = gfp (f.comp g) := - f.dual.map_lfp_comp (OrderHom.dual g) +theorem map_gfp_comp : f (g.comp f).gfp = (f.comp g).gfp := + f.dual.map_lfp_comp g.dual -- Diagonal rule -theorem lfp_lfp (h : α →o α →o α) : lfp (lfp.comp h) = lfp h.onDiag := by - let a := lfp (lfp.comp h) +theorem lfp_lfp (h : α →o α →o α) : (lfp.comp h).lfp = h.onDiag.lfp := by + let a := (lfp.comp h).lfp refine (lfp_le _ ?_).antisymm (lfp_le _ (Eq.le ?_)) · exact lfp_le _ h.onDiag.map_lfp.le have ha : (lfp ∘ h) a = a := (lfp.comp h).map_lfp calc - h a a = h a (lfp (h a)) := congr_arg (h a) ha.symm - _ = lfp (h a) := (h a).map_lfp + h a a = h a (h a).lfp := congr_arg (h a) ha.symm + _ = (h a).lfp := (h a).map_lfp _ = a := ha -theorem gfp_gfp (h : α →o α →o α) : gfp (gfp.comp h) = gfp h.onDiag := - @lfp_lfp αᵒᵈ _ <| (OrderHom.dualIso αᵒᵈ αᵒᵈ).symm.toOrderEmbedding.toOrderHom.comp - (OrderHom.dual h) +theorem gfp_gfp (h : α →o α →o α) : (gfp.comp h).gfp = h.onDiag.gfp := + @lfp_lfp αᵒᵈ _ <| (OrderHom.dualIso αᵒᵈ αᵒᵈ).symm.toOrderEmbedding.toOrderHom.comp h.dual end Eqn @@ -158,25 +154,25 @@ section PrevNext variable [CompleteLattice α] (f : α →o α) -theorem gfp_const_inf_le (x : α) : gfp (const α x ⊓ f) ≤ x := +theorem gfp_const_inf_le (x : α) : (const α x ⊓ f).gfp ≤ x := (gfp_le _) fun _ hb => hb.trans inf_le_left /-- Previous fixed point of a monotone map. If `f` is a monotone self-map of a complete lattice and `x` is a point such that `f x ≤ x`, then `f.prevFixed x hx` is the greatest fixed point of `f` that is less than or equal to `x`. -/ def prevFixed (x : α) (hx : f x ≤ x) : fixedPoints f := - ⟨gfp (const α x ⊓ f), + ⟨(const α x ⊓ f).gfp, calc - f (gfp (const α x ⊓ f)) = x ⊓ f (gfp (const α x ⊓ f)) := + f (const α x ⊓ f).gfp = x ⊓ f (const α x ⊓ f).gfp := Eq.symm <| inf_of_le_right <| (f.mono <| f.gfp_const_inf_le x).trans hx - _ = gfp (const α x ⊓ f) := (const α x ⊓ f).map_gfp + _ = (const α x ⊓ f).gfp := (const α x ⊓ f).map_gfp ⟩ /-- Next fixed point of a monotone map. If `f` is a monotone self-map of a complete lattice and `x` is a point such that `x ≤ f x`, then `f.nextFixed x hx` is the least fixed point of `f` that is greater than or equal to `x`. -/ def nextFixed (x : α) (hx : x ≤ f x) : fixedPoints f := - { f.dual.prevFixed x hx with val := lfp (const α x ⊔ f) } + { f.dual.prevFixed x hx with val := (const α x ⊔ f).lfp } theorem prevFixed_le {x : α} (hx : f x ≤ x) : ↑(f.prevFixed x hx) ≤ x := f.gfp_const_inf_le x @@ -240,7 +236,7 @@ instance : SemilatticeSup (fixedPoints f) := /- porting note: removed `Subtype.partialOrder _` from mathlib3port version, threw `typeclass instance` error and was seemingly unnecessary?-/ instance : SemilatticeInf (fixedPoints f) := - { OrderDual.instSemilatticeInf (fixedPoints (OrderHom.dual f)) with + { OrderDual.instSemilatticeInf (fixedPoints f.dual) with inf := fun x y => f.prevFixed (x ⊓ y) (f.map_inf_fixedPoints_le x y) } -- Porting note: `coe` replaced with `Subtype.val` @@ -271,8 +267,8 @@ instance completeLattice : CompleteLattice (fixedPoints f) where __ := inferInstanceAs (SemilatticeSup (fixedPoints f)) __ := inferInstanceAs (CompleteSemilatticeInf (fixedPoints f)) __ := inferInstanceAs (CompleteSemilatticeSup (fixedPoints f)) - top := ⟨gfp f, f.isFixedPt_gfp⟩ - bot := ⟨lfp f, f.isFixedPt_lfp⟩ + top := ⟨f.gfp, f.isFixedPt_gfp⟩ + bot := ⟨f.lfp, f.isFixedPt_lfp⟩ le_top x := f.le_gfp x.2.ge bot_le x := f.lfp_le x.2.le @@ -281,7 +277,7 @@ open OmegaCompletePartialOrder fixedPoints /-- **Kleene's fixed point Theorem**: The least fixed point in a complete lattice is the supremum of iterating a function on bottom arbitrary often. -/ theorem lfp_eq_sSup_iterate (h : ωScottContinuous f) : - lfp f = ⨆ n, f^[n] ⊥ := by + f.lfp = ⨆ n, f^[n] ⊥ := by apply le_antisymm · apply lfp_le_fixed exact Function.mem_fixedPoints.mp (ωSup_iterate_mem_fixedPoint @@ -290,8 +286,8 @@ theorem lfp_eq_sSup_iterate (h : ωScottContinuous f) : intro a h_a exact ωSup_iterate_le_prefixedPoint ⟨f, h.map_ωSup_of_orderHom⟩ ⊥ bot_le h_a bot_le -theorem gfp_eq_sInf_iterate (h : ωScottContinuous (OrderHom.dual f)) : - gfp f = ⨅ n, f^[n] ⊤ := - lfp_eq_sSup_iterate (OrderHom.dual f) h +theorem gfp_eq_sInf_iterate (h : ωScottContinuous f.dual) : + f.gfp = ⨅ n, f^[n] ⊤ := + lfp_eq_sSup_iterate f.dual h end fixedPoints diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index aaa22efefcb3e..ef9ebbb809b7f 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -488,16 +488,13 @@ protected def dual : (α →o β) ≃ (αᵒᵈ →o βᵒᵈ) where left_inv _ := rfl right_inv _ := rfl --- Porting note: We used to be able to write `(OrderHom.id : α →o α).dual` here rather than --- `OrderHom.dual (OrderHom.id : α →o α)`. --- See https://github.com/leanprover/lean4/issues/1910 @[simp] -theorem dual_id : OrderHom.dual (OrderHom.id : α →o α) = OrderHom.id := +theorem dual_id : (OrderHom.id : α →o α).dual = OrderHom.id := rfl @[simp] theorem dual_comp (g : β →o γ) (f : α →o β) : - OrderHom.dual (g.comp f) = (OrderHom.dual g).comp (OrderHom.dual f) := + (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 00886351437d4..6466c90ddaf43 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -725,18 +725,18 @@ protected def dual : BoundedOrderHom α β ≃ BoundedOrderHom αᵒᵈ βᵒᵈ where - toFun f := ⟨OrderHom.dual f.toOrderHom, f.map_bot', f.map_top'⟩ + toFun f := ⟨f.toOrderHom.dual, f.map_bot', f.map_top'⟩ invFun f := ⟨OrderHom.dual.symm f.toOrderHom, f.map_bot', f.map_top'⟩ left_inv _ := ext fun _ => rfl right_inv _ := ext fun _ => rfl @[simp] -theorem dual_id : BoundedOrderHom.dual (BoundedOrderHom.id α) = BoundedOrderHom.id _ := +theorem dual_id : (BoundedOrderHom.id α).dual = BoundedOrderHom.id _ := rfl @[simp] theorem dual_comp (g : BoundedOrderHom β γ) (f : BoundedOrderHom α β) : - BoundedOrderHom.dual (g.comp f) = g.dual.comp (BoundedOrderHom.dual f) := + (g.comp f).dual = g.dual.comp f.dual := rfl @[simp] diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 9ef57523db1f8..014a33ce871f3 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -153,7 +153,7 @@ theorem map_map (g : β →o γ) (f : α →o β) (a : NonemptyInterval α) : @[simp] theorem dual_map (f : α →o β) (a : NonemptyInterval α) : - dual (a.map f) = a.dual.map (OrderHom.dual f) := + dual (a.map f) = a.dual.map f.dual := rfl /-- Binary pushforward of nonempty intervals. -/ @@ -360,7 +360,7 @@ theorem map_map (g : β →o γ) (f : α →o β) (s : Interval α) : (s.map f). Option.map_map _ _ _ @[simp] -theorem dual_map (f : α →o β) (s : Interval α) : dual (s.map f) = s.dual.map (OrderHom.dual f) := by +theorem dual_map (f : α →o β) (s : Interval α) : dual (s.map f) = s.dual.map f.dual := by cases s · rfl · exact WithBot.map_comm rfl _ diff --git a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean index a5c9f5c018629..02acdbd1461b6 100644 --- a/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean +++ b/Mathlib/SetTheory/Cardinal/SchroederBernstein.lean @@ -49,8 +49,7 @@ theorem schroeder_bernstein {f : α → β} {g : β → α} (hf : Function.Injec { toFun := fun s => (g '' (f '' s)ᶜ)ᶜ monotone' := fun s t hst => compl_subset_compl.mpr <| image_subset _ <| compl_subset_compl.mpr <| image_subset _ hst } - -- Porting note: dot notation `F.lfp` doesn't work here - set s : Set α := OrderHom.lfp F + set s : Set α := F.lfp have hs : (g '' (f '' s)ᶜ)ᶜ = s := F.map_lfp have hns : g '' (f '' s)ᶜ = sᶜ := compl_injective (by simp [hs]) set g' := invFun g diff --git a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean index 97272de665f4f..4d1e937f10243 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPointApproximants.lean @@ -226,9 +226,9 @@ theorem lfpApprox_le_of_mem_fixedPoints {a : α} /-- The approximation sequence converges at the successor of the domain's cardinality to the least fixed point if starting from `⊥` -/ -theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = lfp f := by +theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = f.lfp := by apply le_antisymm - · have h_lfp : ∃ y : fixedPoints f, lfp f = y := by use ⊥; exact rfl + · have h_lfp : ∃ y : fixedPoints f, f.lfp = y := by use ⊥; exact rfl let ⟨y, h_y⟩ := h_lfp; rw [h_y] exact lfpApprox_le_of_mem_fixedPoints f ⊥ y.2 bot_le (ord <| succ #α) · have h_fix : ∃ y : fixedPoints f, lfpApprox f ⊥ (ord <| succ #α) = y := by @@ -238,7 +238,7 @@ theorem lfpApprox_ord_eq_lfp : lfpApprox f ⊥ (ord <| succ #α) = lfp f := by exact lfp_le_fixed f x.prop /-- Some approximation of the least fixed point starting from `⊥` is the least fixed point. -/ -theorem lfp_mem_range_lfpApprox : lfp f ∈ Set.range (lfpApprox f ⊥) := by +theorem lfp_mem_range_lfpApprox : f.lfp ∈ Set.range (lfpApprox f ⊥) := by use ord <| succ #α exact lfpApprox_ord_eq_lfp f @@ -256,52 +256,52 @@ decreasing_by exact h unseal gfpApprox lfpApprox theorem gfpApprox_antitone : Antitone (gfpApprox f x) := - lfpApprox_monotone (OrderHom.dual f) x + lfpApprox_monotone f.dual x theorem gfpApprox_le {a : Ordinal} : gfpApprox f x a ≤ x := - le_lfpApprox (OrderHom.dual f) x + le_lfpApprox f.dual x theorem gfpApprox_add_one (h : f x ≤ x) (a : Ordinal) : gfpApprox f x (a+1) = f (gfpApprox f x a) := - lfpApprox_add_one (OrderHom.dual f) x h a + lfpApprox_add_one f.dual x h a theorem gfpApprox_mono_left : Monotone (gfpApprox : (α →o α) → _) := by intro f g h - have : OrderHom.dual g ≤ OrderHom.dual f := h + have : g.dual ≤ f.dual := h exact lfpApprox_mono_left this theorem gfpApprox_mono_mid : Monotone (gfpApprox f) := - fun _ _ h => lfpApprox_mono_mid (OrderHom.dual f) h + fun _ _ h => lfpApprox_mono_mid f.dual h /-- The approximations of the greatest fixed point stabilize at a fixed point of `f` -/ theorem gfpApprox_eq_of_mem_fixedPoints {a b : Ordinal} (h_init : f x ≤ x) (h_ab : a ≤ b) (h : gfpApprox f x a ∈ fixedPoints f) : gfpApprox f x b = gfpApprox f x a := - lfpApprox_eq_of_mem_fixedPoints (OrderHom.dual f) x h_init h_ab h + lfpApprox_eq_of_mem_fixedPoints f.dual x h_init h_ab h /-- There are distinct indices smaller than the successor of the domain's cardinality yielding the same value -/ theorem exists_gfpApprox_eq_gfpApprox : ∃ a < ord <| succ #α, ∃ b < ord <| succ #α, a ≠ b ∧ gfpApprox f x a = gfpApprox f x b := - exists_lfpApprox_eq_lfpApprox (OrderHom.dual f) x + exists_lfpApprox_eq_lfpApprox f.dual x /-- The approximation at the index of the successor of the domain's cardinality is a fixed point -/ lemma gfpApprox_ord_mem_fixedPoint (h_init : f x ≤ x) : gfpApprox f x (ord <| succ #α) ∈ fixedPoints f := - lfpApprox_ord_mem_fixedPoint (OrderHom.dual f) x h_init + lfpApprox_ord_mem_fixedPoint f.dual x h_init /-- Every value of the approximation is greater or equal than every fixed point of `f` less or equal than the initial value -/ lemma le_gfpApprox_of_mem_fixedPoints {a : α} (h_a : a ∈ fixedPoints f) (h_le_init : a ≤ x) (i : Ordinal) : a ≤ gfpApprox f x i := - lfpApprox_le_of_mem_fixedPoints (OrderHom.dual f) x h_a h_le_init i + lfpApprox_le_of_mem_fixedPoints f.dual x h_a h_le_init i /-- The approximation sequence converges at the successor of the domain's cardinality to the greatest fixed point if starting from `⊥` -/ -theorem gfpApprox_ord_eq_gfp : gfpApprox f ⊤ (ord <| succ #α) = gfp f := - lfpApprox_ord_eq_lfp (OrderHom.dual f) +theorem gfpApprox_ord_eq_gfp : gfpApprox f ⊤ (ord <| succ #α) = f.gfp := + lfpApprox_ord_eq_lfp f.dual /-- Some approximation of the least fixed point starting from `⊤` is the greatest fixed point. -/ -theorem gfp_mem_range_gfpApprox : gfp f ∈ Set.range (gfpApprox f ⊤) := - lfp_mem_range_lfpApprox (OrderHom.dual f) +theorem gfp_mem_range_gfpApprox : f.gfp ∈ Set.range (gfpApprox f ⊤) := + lfp_mem_range_lfpApprox f.dual end OrdinalApprox From f700f821a8e256bf132e1e9fb6c433f83a5ceda4 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Sat, 23 Nov 2024 06:30:16 +0000 Subject: [PATCH 054/250] feat(FieldTheory): minimal polynomials determine an algebraic extension (Isaacs 1980) (#18412) If E/F is an algebraic extension and K/F is an arbitrary extension, and if the minimal polynomial of every element of E over F has a root in K, we show there exists an F-embedding from E into K. If E/F and K/F have the same set of minimal polynomials, they are then isomorphic. Another corollary is that an algebraic extension E/F is an algebraic closure if every (monic irreducible) polynomial in F[X] has a root in E. These results are proven in FieldTheory/Isaacs. An important definition (`IsExtendible`) and two important lemmas towards these theorems concern `IntermediateField.Lifts` and reside in FieldTheory/Extension. The lemmas say that the supremum of a chain of extendible lifts is also extendible, and that a lift with full domain exists if the bottom lift is extendible. We do slight refactoring to extract `Lifts.union`, and provide new API lemmas for Lifts. Co-authored-by: Junyan Xu --- Mathlib.lean | 1 + Mathlib/FieldTheory/Extension.lean | 162 +++++++++++++++++++-- Mathlib/FieldTheory/IsAlgClosed/Basic.lean | 8 +- Mathlib/FieldTheory/Isaacs.lean | 115 +++++++++++++++ docs/references.bib | 13 ++ 5 files changed, 280 insertions(+), 19 deletions(-) create mode 100644 Mathlib/FieldTheory/Isaacs.lean diff --git a/Mathlib.lean b/Mathlib.lean index d142e08e33749..075b258277b1b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2951,6 +2951,7 @@ import Mathlib.FieldTheory.IsAlgClosed.Classification import Mathlib.FieldTheory.IsAlgClosed.Spectrum import Mathlib.FieldTheory.IsPerfectClosure import Mathlib.FieldTheory.IsSepClosed +import Mathlib.FieldTheory.Isaacs import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension diff --git a/Mathlib/FieldTheory/Extension.lean b/Mathlib/FieldTheory/Extension.lean index da3e16176910d..142e5ed48085b 100644 --- a/Mathlib/FieldTheory/Extension.lean +++ b/Mathlib/FieldTheory/Extension.lean @@ -1,8 +1,9 @@ /- Copyright (c) 2020 Thomas Browning. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Thomas Browning +Authors: Thomas Browning, Junyan Xu -/ +import Mathlib.Data.Fintype.Order import Mathlib.FieldTheory.Adjoin /-! @@ -12,6 +13,12 @@ import Mathlib.FieldTheory.Adjoin field extensions, K is another extension of F, and `f` is an embedding of L/F into K/F, such that the minimal polynomials of a set of generators of E/L splits in K (via `f`), then `f` extends to an embedding of E/F into K/F. + +## Reference + +[Isaacs1980] *Roots of Polynomials in Algebraic Extensions of Fields*, +The American Mathematical Monthly + -/ open Polynomial @@ -29,6 +36,8 @@ structure Lifts where variable {F E K} +namespace Lifts + instance : PartialOrder (Lifts F E K) where le L₁ L₂ := ∃ h : L₁.carrier ≤ L₂.carrier, ∀ x, L₂.emb (inclusion h x) = L₁.emb x le_refl L := ⟨le_rfl, by simp⟩ @@ -52,25 +61,144 @@ noncomputable instance : OrderBot (Lifts F E K) where noncomputable instance : Inhabited (Lifts F E K) := ⟨⊥⟩ -/-- A chain of lifts has an upper bound. -/ -theorem Lifts.exists_upper_bound (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) : - ∃ ub, ∀ a ∈ c, a ≤ ub := by +variable {L₁ L₂ : Lifts F E K} + +theorem le_iff : L₁ ≤ L₂ ↔ + ∃ h : L₁.carrier ≤ L₂.carrier, L₂.emb.comp (inclusion h) = L₁.emb := by + simp_rw [AlgHom.ext_iff]; rfl + +theorem eq_iff_le_carrier_eq : L₁ = L₂ ↔ L₁ ≤ L₂ ∧ L₁.carrier = L₂.carrier := + ⟨fun eq ↦ ⟨eq.le, congr_arg _ eq⟩, fun ⟨le, eq⟩ ↦ le.antisymm ⟨eq.ge, fun x ↦ (le.2 ⟨x, _⟩).symm⟩⟩ + +theorem eq_iff : L₁ = L₂ ↔ + ∃ h : L₁.carrier = L₂.carrier, L₂.emb.comp (inclusion h.le) = L₁.emb := by + rw [eq_iff_le_carrier_eq, le_iff] + exact ⟨fun h ↦ ⟨h.2, h.1.2⟩, fun h ↦ ⟨⟨h.1.le, h.2⟩, h.1⟩⟩ + +theorem lt_iff_le_carrier_ne : L₁ < L₂ ↔ L₁ ≤ L₂ ∧ L₁.carrier ≠ L₂.carrier := by + rw [lt_iff_le_and_ne, and_congr_right]; intro h; simp_rw [Ne, eq_iff_le_carrier_eq, h, true_and] + +theorem lt_iff : L₁ < L₂ ↔ + ∃ h : L₁.carrier < L₂.carrier, L₂.emb.comp (inclusion h.le) = L₁.emb := by + rw [lt_iff_le_carrier_ne, le_iff] + exact ⟨fun h ↦ ⟨h.1.1.lt_of_ne h.2, h.1.2⟩, fun h ↦ ⟨⟨h.1.le, h.2⟩, h.1.ne⟩⟩ + +theorem le_of_carrier_le_iSup {ι} {ρ : ι → Lifts F E K} {σ τ : Lifts F E K} + (hσ : ∀ i, ρ i ≤ σ) (hτ : ∀ i, ρ i ≤ τ) (carrier_le : σ.carrier ≤ ⨆ i, (ρ i).carrier) : + σ ≤ τ := + le_iff.mpr ⟨carrier_le.trans (iSup_le fun i ↦ (hτ i).1), algHom_ext_of_eq_adjoin _ + (carrier_le.antisymm (iSup_le fun i ↦ (hσ i).1)|>.trans <| iSup_eq_adjoin _ _) fun x hx ↦ + have ⟨i, hx⟩ := Set.mem_iUnion.mp hx + ((hτ i).2 ⟨x, hx⟩).trans ((hσ i).2 ⟨x, hx⟩).symm⟩ + +/-- `σ : L →ₐ[F] K` is an extendible lift ("extendible pair" in [Isaacs1980]) if for every +intermediate field `M` that is finite-dimensional over `L`, `σ` extends to some `M →ₐ[F] K`. +In our definition we only require `M` to be finitely generated over `L`, which is equivalent +if the ambient field `E` is algebraic over `F` (which is the case in our main application). +We also allow the domain of the extension to be an intermediate field that properly contains `M`, +since one can always restrict the domain to `M`. -/ +def IsExtendible (σ : Lifts F E K) : Prop := + ∀ S : Finset E, ∃ τ ≥ σ, (S : Set E) ⊆ τ.carrier + +section Chain +variable (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) + +/-- The union of a chain of lifts. -/ +noncomputable def union : Lifts F E K := let t (i : ↑(insert ⊥ c)) := i.val.carrier - let t' (i) := (t i).toSubalgebra have hc := hc.insert fun _ _ _ ↦ .inl bot_le have dir : Directed (· ≤ ·) t := hc.directedOn.directed_val.mono_comp _ fun _ _ h ↦ h.1 - refine ⟨⟨iSup t, (Subalgebra.iSupLift t' dir (fun i ↦ i.val.emb) (fun i j h ↦ ?_) _ rfl).comp - (Subalgebra.equivOfEq _ _ <| toSubalgebra_iSup_of_directed dir)⟩, - fun L hL ↦ have hL := Set.mem_insert_of_mem ⊥ hL; ⟨le_iSup t ⟨L, hL⟩, fun x ↦ ?_⟩⟩ - · refine AlgHom.ext fun x ↦ (hc.total i.2 j.2).elim (fun hij ↦ (hij.snd x).symm) fun hji ↦ ?_ - erw [AlgHom.comp_apply, ← hji.snd (Subalgebra.inclusion h x), - inclusion_inclusion, inclusion_self, AlgHom.id_apply x] - · dsimp only [AlgHom.comp_apply] - exact Subalgebra.iSupLift_inclusion (K := t') (i := ⟨L, hL⟩) x (le_iSup t' ⟨L, hL⟩) + ⟨iSup t, (Subalgebra.iSupLift (toSubalgebra <| t ·) dir (·.val.emb) (fun i j h ↦ + AlgHom.ext fun x ↦ (hc.total i.2 j.2).elim (fun hij ↦ (hij.snd x).symm) fun hji ↦ by + erw [AlgHom.comp_apply, ← hji.snd (Subalgebra.inclusion h x), + inclusion_inclusion, inclusion_self, AlgHom.id_apply x]) _ rfl).comp + (Subalgebra.equivOfEq _ _ <| toSubalgebra_iSup_of_directed dir)⟩ + +theorem le_union ⦃σ : Lifts F E K⦄ (hσ : σ ∈ c) : σ ≤ union c hc := + have hσ := Set.mem_insert_of_mem ⊥ hσ + let t (i : ↑(insert ⊥ c)) := i.val.carrier + ⟨le_iSup t ⟨σ, hσ⟩, fun x ↦ by + dsimp only [union, AlgHom.comp_apply] + exact Subalgebra.iSupLift_inclusion (K := (toSubalgebra <| t ·)) + (i := ⟨σ, hσ⟩) x (le_iSup (toSubalgebra <| t ·) ⟨σ, hσ⟩)⟩ + +theorem carrier_union : (union c hc).carrier = ⨆ i : c, i.1.carrier := + le_antisymm (iSup_le <| by rintro ⟨i, rfl|hi⟩; exacts [bot_le, le_iSup_of_le ⟨i, hi⟩ le_rfl]) <| + iSup_le fun i ↦ le_iSup_of_le ⟨i, .inr i.2⟩ le_rfl + +/-- A chain of lifts has an upper bound. -/ +theorem exists_upper_bound (c : Set (Lifts F E K)) (hc : IsChain (· ≤ ·) c) : + ∃ ub, ∀ a ∈ c, a ≤ ub := ⟨_, le_union c hc⟩ + +theorem union_isExtendible [alg : Algebra.IsAlgebraic F E] + [Nonempty c] (hext : ∀ σ ∈ c, σ.IsExtendible) : + (union c hc).IsExtendible := fun S ↦ by + let Ω := adjoin F (S : Set E) →ₐ[F] K + have ⟨ω, hω⟩ : ∃ ω : Ω, ∀ π : c, ∃ θ ≥ π.1, ⟨_, ω⟩ ≤ θ ∧ θ.carrier = π.1.1 ⊔ adjoin F S := by + by_contra!; choose π hπ using this + have := finiteDimensional_adjoin (S := (S : Set E)) fun _ _ ↦ (alg.isIntegral).1 _ + have ⟨π₀, hπ₀⟩ := hc.directed.finite_le π + have ⟨θ, hθπ, hθ⟩ := hext _ π₀.2 S + rw [← adjoin_le_iff] at hθ + let θ₀ := θ.emb.comp (inclusion hθ) + have := (hπ₀ θ₀).trans hθπ + exact hπ θ₀ ⟨_, θ.emb.comp <| inclusion <| sup_le this.1 hθ⟩ + ⟨le_sup_left, this.2⟩ ⟨le_sup_right, fun _ ↦ rfl⟩ rfl + choose θ ge hθ eq using hω + have : IsChain (· ≤ ·) (Set.range θ) := by + simp_rw [← restrictScalars_adjoin_eq_sup, restrictScalars_adjoin] at eq + rintro _ ⟨π₁, rfl⟩ _ ⟨π₂, rfl⟩ - + wlog h : π₁ ≤ π₂ generalizing π₁ π₂ + · exact (this _ _ <| (hc.total π₁.2 π₂.2).resolve_left h).symm + refine .inl (le_iff.mpr ⟨?_, algHom_ext_of_eq_adjoin _ (eq _) ?_⟩) + · rw [eq, eq]; exact adjoin.mono _ _ _ (Set.union_subset_union_left _ h.1) + rintro x (hx|hx) + · change (θ π₂).emb (inclusion (ge π₂).1 <| inclusion h.1 ⟨x, hx⟩) = + (θ π₁).emb (inclusion (ge π₁).1 ⟨x, hx⟩) + rw [(ge π₁).2, (ge π₂).2, h.2] + · change (θ π₂).emb (inclusion (hθ π₂).1 ⟨x, subset_adjoin _ _ hx⟩) = + (θ π₁).emb (inclusion (hθ π₁).1 ⟨x, subset_adjoin _ _ hx⟩) + rw [(hθ π₁).2, (hθ π₂).2] + refine ⟨union _ this, le_of_carrier_le_iSup (fun π ↦ le_union c hc π.2) + (fun π ↦ (ge π).trans <| le_union _ _ ⟨_, rfl⟩) (carrier_union _ _).le, ?_⟩ + simp_rw [carrier_union, iSup_range', eq] + exact (subset_adjoin _ _).trans (SetLike.coe_subset_coe.mpr <| + le_sup_right.trans <| le_iSup_of_le (Classical.arbitrary _) le_rfl) + +end Chain + +theorem nonempty_algHom_of_exist_lifts_finset [alg : Algebra.IsAlgebraic F E] + (h : ∀ S : Finset E, ∃ σ : Lifts F E K, (S : Set E) ⊆ σ.carrier) : + Nonempty (E →ₐ[F] K) := by + have : (⊥ : Lifts F E K).IsExtendible := fun S ↦ have ⟨σ, hσ⟩ := h S; ⟨σ, bot_le, hσ⟩ + have ⟨ϕ, hϕ⟩ := zorn_le₀ {ϕ : Lifts F E K | ϕ.IsExtendible} + fun c hext hc ↦ (isEmpty_or_nonempty c).elim + (fun _ ↦ ⟨⊥, this, fun ϕ hϕ ↦ isEmptyElim (⟨ϕ, hϕ⟩ : c)⟩) + fun _ ↦ ⟨_, union_isExtendible c hc hext, le_union c hc⟩ + suffices ϕ.carrier = ⊤ from ⟨ϕ.emb.comp <| ((equivOfEq this).trans topEquiv).symm⟩ + by_contra! + obtain ⟨α, -, hα⟩ := SetLike.exists_of_lt this.lt_top + let _ : Algebra ϕ.carrier K := ϕ.emb.toAlgebra + let Λ := ϕ.carrier⟮α⟯ →ₐ[ϕ.carrier] K + have := finiteDimensional_adjoin (S := {α}) fun _ _ ↦ ((alg.tower_top ϕ.carrier).isIntegral).1 _ + let L (σ : Λ) : Lifts F E K := ⟨ϕ.carrier⟮α⟯.restrictScalars F, σ.restrictScalars F⟩ + have hL (σ : Λ) : ϕ < L σ := lt_iff.mpr + ⟨by simpa only [restrictScalars_adjoin_eq_sup, left_lt_sup, adjoin_simple_le_iff], + AlgHom.coe_ringHom_injective σ.comp_algebraMap⟩ + have ⟨(ϕ_ext : ϕ.IsExtendible), ϕ_max⟩ := maximal_iff_forall_gt.mp hϕ + simp_rw [Set.mem_setOf, IsExtendible] at ϕ_max; push_neg at ϕ_max + choose S hS using fun σ : Λ ↦ ϕ_max (hL σ) + classical + have ⟨θ, hθϕ, hθ⟩ := ϕ_ext ({α} ∪ Finset.univ.biUnion S) + simp_rw [Finset.coe_union, Set.union_subset_iff, Finset.coe_singleton, Set.singleton_subset_iff, + Finset.coe_biUnion, Finset.coe_univ, Set.mem_univ, Set.iUnion_true, Set.iUnion_subset_iff] at hθ + have : ϕ.carrier⟮α⟯.restrictScalars F ≤ θ.carrier := by + rw [restrictScalars_adjoin_eq_sup, sup_le_iff, adjoin_simple_le_iff]; exact ⟨hθϕ.1, hθ.1⟩ + exact hS ⟨(θ.emb.comp <| inclusion this).toRingHom, hθϕ.2⟩ θ ⟨this, fun _ ↦ rfl⟩ (hθ.2 _) /-- Given a lift `x` and an integral element `s : E` over `x.carrier` whose conjugates over `x.carrier` are all in `K`, we can extend the lift to a lift whose carrier contains `s`. -/ -theorem Lifts.exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral x.carrier s) +theorem exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral x.carrier s) (h2 : (minpoly x.carrier s).Splits x.emb.toRingHom) : ∃ y, x ≤ y ∧ s ∈ y.carrier := have I2 := (minpoly.degree_pos h1).ne' letI : Algebra x.carrier K := x.emb.toRingHom.toAlgebra @@ -87,11 +215,13 @@ theorem Lifts.exists_lift_of_splits' (x : Lifts F E K) {s : E} (h1 : IsIntegral /-- Given an integral element `s : E` over `F` whose `F`-conjugates are all in `K`, any lift can be extended to one whose carrier contains `s`. -/ -theorem Lifts.exists_lift_of_splits (x : Lifts F E K) {s : E} (h1 : IsIntegral F s) +theorem exists_lift_of_splits (x : Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : (minpoly F s).Splits (algebraMap F K)) : ∃ y, x ≤ y ∧ s ∈ y.carrier := - Lifts.exists_lift_of_splits' x h1.tower_top <| h1.minpoly_splits_tower_top' <| by + exists_lift_of_splits' x h1.tower_top <| h1.minpoly_splits_tower_top' <| by rwa [← x.emb.comp_algebraMap] at h2 +end Lifts + section private theorem exists_algHom_adjoin_of_splits'' {L : IntermediateField F E} diff --git a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean index c2bc06cd1f067..c164bdbac9521 100644 --- a/Mathlib/FieldTheory/IsAlgClosed/Basic.lean +++ b/Mathlib/FieldTheory/IsAlgClosed/Basic.lean @@ -32,10 +32,12 @@ polynomial in `k` splits. algebraic closure, algebraically closed -## TODO +## Main reults -- Prove that if `K / k` is algebraic, and any monic irreducible polynomial over `k` has a root - in `K`, then `K` is algebraically closed (in fact an algebraic closure of `k`). +- `IsAlgClosure.of_splits`: if `K / k` is algebraic, and every monic irreducible polynomial over + `k` splits in `K`, then `K` is algebraically closed (in fact an algebraic closure of `k`). + For the stronger fact that only requires every such polynomial has a root in `K`, + see `IsAlgClosure.of_exist_roots`. Reference: , Theorem 2 diff --git a/Mathlib/FieldTheory/Isaacs.lean b/Mathlib/FieldTheory/Isaacs.lean new file mode 100644 index 0000000000000..2219c66b3c027 --- /dev/null +++ b/Mathlib/FieldTheory/Isaacs.lean @@ -0,0 +1,115 @@ +/- +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.PrimitiveElement +import Mathlib.GroupTheory.CosetCover + +/-! +# Algebraic extensions are determined by their sets of minimal polynomials up to isomorphism + +## Main results + +`Field.nonempty_algHom_of_exist_roots` says if `E/F` and `K/F` are field extensions +with `E/F` algebraic, and if the minimal polynomial of every element of `E` over `F` has a root +in `K`, then there exists an `F`-embedding of `E` into `K`. If `E/F` and `K/F` have the same +set of minimal polynomials, then `E` and `K` are isomorphic as `F`-algebras. As a corollary: + +`IsAlgClosure.of_exist_roots`: if `E/F` is algebraic and every monic irreducible polynomial +in `F[X]` has a root in `E`, then `E` is an algebraic closure of `F`. + +## Reference + +[Isaacs1980] *Roots of Polynomials in Algebraic Extensions of Fields*, +The American Mathematical Monthly + +-/ + +namespace Field + +open Polynomial IntermediateField + +variable {F E K : Type*} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] +variable [alg : Algebra.IsAlgebraic F E] + +theorem nonempty_algHom_of_exist_roots (h : ∀ x : E, ∃ y : K, aeval y (minpoly F x) = 0) : + Nonempty (E →ₐ[F] K) := by + refine Lifts.nonempty_algHom_of_exist_lifts_finset fun S ↦ ⟨⟨adjoin F S, ?_⟩, subset_adjoin _ _⟩ + let p := (S.prod <| minpoly F).map (algebraMap F K) + let K' := SplittingField p + have splits s (hs : s ∈ S) : (minpoly F s).Splits (algebraMap F K') := by + apply splits_of_splits_of_dvd _ + (Finset.prod_ne_zero_iff.mpr fun _ _ ↦ minpoly.ne_zero <| (alg.isIntegral).1 _) + ((splits_map_iff _ _).mp <| SplittingField.splits p) (Finset.dvd_prod_of_mem _ hs) + let K₀ := (⊥ : IntermediateField K K').restrictScalars F + let FS := adjoin F (S : Set E) + let Ω := FS →ₐ[F] K' + have := finiteDimensional_adjoin (S := (S : Set E)) fun _ _ ↦ (alg.isIntegral).1 _ + let M (ω : Ω) := Subalgebra.toSubmodule (K₀.comap ω).toSubalgebra + have : ⋃ ω : Ω, (M ω : Set FS) = Set.univ := + Set.eq_univ_of_forall fun ⟨α, hα⟩ ↦ Set.mem_iUnion.mpr <| by + have ⟨β, hβ⟩ := h α + let ϕ : F⟮α⟯ →ₐ[F] K' := (IsScalarTower.toAlgHom _ _ _).comp ((AdjoinRoot.liftHom _ _ hβ).comp + (adjoinRootEquivAdjoin F <| (alg.isIntegral).1 _).symm.toAlgHom) + have ⟨ω, hω⟩ := exists_algHom_adjoin_of_splits + (fun s hs ↦ ⟨(alg.isIntegral).1 _, splits s hs⟩) ϕ (adjoin_simple_le_iff.mpr hα) + refine ⟨ω, β, ((DFunLike.congr_fun hω <| AdjoinSimple.gen F α).trans ?_).symm⟩ + rw [AlgHom.comp_apply, AlgHom.comp_apply, AlgEquiv.coe_algHom, + adjoinRootEquivAdjoin_symm_apply_gen, AdjoinRoot.liftHom_root] + rfl + have ω : ∃ ω : Ω, ⊤ ≤ M ω := by + cases finite_or_infinite F + · have ⟨α, hα⟩ := exists_primitive_element_of_finite_bot F FS + have ⟨ω, hω⟩ := Set.mem_iUnion.mp (this ▸ Set.mem_univ α) + exact ⟨ω, show ⊤ ≤ K₀.comap ω by rwa [← hα, adjoin_simple_le_iff]⟩ + · simp_rw [top_le_iff, Subspace.exists_eq_top_of_iUnion_eq_univ this] + exact ((botEquiv K K').toAlgHom.restrictScalars F).comp + (ω.choose.codRestrict K₀.toSubalgebra fun x ↦ ω.choose_spec trivial) + +theorem nonempty_algHom_of_minpoly_eq + (h : ∀ x : E, ∃ y : K, minpoly F x = minpoly F y) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_exist_roots fun x ↦ have ⟨y, hy⟩ := h x; ⟨y, by rw [hy, minpoly.aeval]⟩ + +theorem nonempty_algHom_of_range_minpoly_subset + (h : Set.range (@minpoly F E _ _ _) ⊆ Set.range (@minpoly F K _ _ _)) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_minpoly_eq fun x ↦ have ⟨y, hy⟩ := h ⟨x, rfl⟩; ⟨y, hy.symm⟩ + +theorem nonempty_algEquiv_of_range_minpoly_eq + (h : Set.range (@minpoly F E _ _ _) = Set.range (@minpoly F K _ _ _)) : + Nonempty (E ≃ₐ[F] K) := + have ⟨σ⟩ := nonempty_algHom_of_range_minpoly_subset h.le + have : Algebra.IsAlgebraic F K := ⟨fun y ↦ IsIntegral.isAlgebraic <| by + by_contra hy + have ⟨x, hx⟩ := h.ge ⟨y, rfl⟩ + rw [minpoly.eq_zero hy] at hx + exact minpoly.ne_zero ((alg.isIntegral).1 x) hx⟩ + have ⟨τ⟩ := nonempty_algHom_of_range_minpoly_subset h.ge + ⟨.ofBijective _ (Algebra.IsAlgebraic.algHom_bijective₂ σ τ).1⟩ + +theorem nonempty_algHom_of_aeval_eq_zero_subset + (h : {p : F[X] | ∃ x : E, aeval x p = 0} ⊆ {p | ∃ y : K, aeval y p = 0}) : + Nonempty (E →ₐ[F] K) := + nonempty_algHom_of_minpoly_eq fun x ↦ + have ⟨y, hy⟩ := h ⟨_, minpoly.aeval F x⟩ + ⟨y, (minpoly.eq_iff_aeval_minpoly_eq_zero <| (alg.isIntegral).1 x).mpr hy⟩ + +theorem nonempty_algEquiv_of_aeval_eq_zero_eq [Algebra.IsAlgebraic F K] + (h : {p : F[X] | ∃ x : E, aeval x p = 0} = {p | ∃ y : K, aeval y p = 0}) : + Nonempty (E ≃ₐ[F] K) := + have ⟨σ⟩ := nonempty_algHom_of_aeval_eq_zero_subset h.le + have ⟨τ⟩ := nonempty_algHom_of_aeval_eq_zero_subset h.ge + ⟨.ofBijective _ (Algebra.IsAlgebraic.algHom_bijective₂ σ τ).1⟩ + +theorem _root_.IsAlgClosure.of_exist_roots + (h : ∀ p : F[X], p.Monic → Irreducible p → ∃ x : E, aeval x p = 0) : + IsAlgClosure F E := + .of_splits fun p _ _ ↦ + have ⟨σ⟩ := nonempty_algHom_of_exist_roots fun x : p.SplittingField ↦ + have := Algebra.IsAlgebraic.isIntegral (K := F).1 x + h _ (minpoly.monic this) (minpoly.irreducible this) + splits_of_algHom (SplittingField.splits _) σ + +end Field diff --git a/docs/references.bib b/docs/references.bib index 392120a19de1e..f227b4270fd17 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -2083,6 +2083,19 @@ @Book{ IrelandRosen1990 url = {https://doi.org/10.1007/978-1-4757-2103-4} } +@Article{ Isaacs1980, + author = {Isaacs, I. M.}, + title = {Roots of Polynomials in Algebraic Extensions of Fields}, + publisher = {Taylor \& Francis}, + year = {1980}, + pages = {543--544}, + journal = {The American Mathematical Monthly}, + volume = {87}, + number = {7}, + doi = {10.1080/00029890.1980.11995085}, + url = {https://doi.org/10.1080/00029890.1980.11995085} +} + @Book{ iversen, title = {Generic Local Structure of the Morphisms in Commutative Algebra}, From 1ef6e445ac54d2bb1d2568e3bc000b8512a0a765 Mon Sep 17 00:00:00 2001 From: Hannah Scholz Date: Sat, 23 Nov 2024 06:30:17 +0000 Subject: [PATCH 055/250] feat: image of any interval under `affineHomeomorph` (#19301) --- Mathlib/Data/Set/Pointwise/Interval.lean | 37 ++++++++++++++++++++++++ Mathlib/Topology/Algebra/Field.lean | 21 ++++++++++++++ 2 files changed, 58 insertions(+) diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Data/Set/Pointwise/Interval.lean index 535b6e9100c34..e6846ddbc75b3 100644 --- a/Mathlib/Data/Set/Pointwise/Interval.lean +++ b/Mathlib/Data/Set/Pointwise/Interval.lean @@ -725,6 +725,22 @@ theorem image_mul_left_Ioo {a : α} (h : 0 < a) (b c : α) : (a * ·) '' Ioo b c = Ioo (a * b) (a * c) := by convert image_mul_right_Ioo b c h using 1 <;> simp only [mul_comm _ a] +theorem image_mul_right_Ico (a b : α) {c : α} (h : 0 < c) : + (fun x => x * c) '' Ico a b = Ico (a * c) (b * c) := + ((Units.mk0 c h.ne').mulRight.image_eq_preimage _).trans (by simp [h, division_def]) + +theorem image_mul_left_Ico {a : α} (h : 0 < a) (b c : α) : + (a * ·) '' Ico b c = Ico (a * b) (a * c) := by + convert image_mul_right_Ico b c h using 1 <;> simp only [mul_comm _ a] + +theorem image_mul_right_Ioc (a b : α) {c : α} (h : 0 < c) : + (fun x => x * c) '' Ioc a b = Ioc (a * c) (b * c) := + ((Units.mk0 c h.ne').mulRight.image_eq_preimage _).trans (by simp [h, division_def]) + +theorem image_mul_left_Ioc {a : α} (h : 0 < a) (b c : α) : + (a * ·) '' Ioc b c = Ioc (a * b) (a * c) := by + convert image_mul_right_Ioc b c h using 1 <;> simp only [mul_comm _ a] + /-- The (pre)image under `inv` of `Ioo 0 a` is `Ioi a⁻¹`. -/ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by ext x @@ -754,6 +770,27 @@ theorem image_affine_Icc' {a : α} (h : 0 < a) (b c d : α) : rwa [Set.image_image] at this rw [image_mul_left_Icc' h, image_add_const_Icc] +@[simp] +theorem image_affine_Ico {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ico c d = Ico (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ico c d) = Ico (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ico h, image_add_const_Ico] + +@[simp] +theorem image_affine_Ioc {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ioc c d = Ioc (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ioc c d) = Ioc (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ioc h, image_add_const_Ioc] + +@[simp] +theorem image_affine_Ioo {a : α} (h : 0 < a) (b c d : α) : + (a * · + b) '' Ioo c d = Ioo (a * c + b) (a * d + b) := by + suffices (· + b) '' ((a * ·) '' Ioo c d) = Ioo (a * c + b) (a * d + b) by + rwa [Set.image_image] at this + rw [image_mul_left_Ioo h, image_add_const_Ioo] + end LinearOrderedField end Set diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 48b1183bc9f01..7e78ef983d7f6 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Order.LocalExtr +import Mathlib.Data.Set.Pointwise.Interval /-! # Topological fields @@ -91,6 +92,26 @@ def affineHomeomorph (a b : 𝕜) (h : a ≠ 0) : 𝕜 ≃ₜ 𝕜 where exact mul_div_cancel_left₀ x h right_inv y := by simp [mul_div_cancel₀ _ h] +theorem affineHomeomorph_image_Icc {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Icc c d = Set.Icc (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ico {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ico c d = Set.Ico (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ioc {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ioc c d = Set.Ioc (a * c + b) (a * d + b) := by + simp [h] + +theorem affineHomeomorph_image_Ioo {𝕜 : Type*} [LinearOrderedField 𝕜] [TopologicalSpace 𝕜] + [TopologicalRing 𝕜] (a b c d : 𝕜) (h : 0 < a) : + affineHomeomorph a b h.ne' '' Set.Ioo c d = Set.Ioo (a * c + b) (a * d + b) := by + simp [h] + end affineHomeomorph section LocalExtr From 02990e46811c8ead901b6c056cf91dc7e3e80c10 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 06:30:18 +0000 Subject: [PATCH 056/250] chore(Algebra/Module/Torsion): change adaptation note links to lean4 issue #1910 to refer to lean4 issue #1629 instead (#19381) These notes refer to the fact that `LinearMap.ker` and friends cannot be used as `f.ker`. However, this fails not because the function `LinearMap.ker` needs to be coerced to a proper function (which the issue leanprover/lean4#1910 is about), but because the argument `f` has a type (`LinearMap`) that doesn't occur directly as type of an argument to `LinearMap.ker`. The issue leanprover/lean4#1629 *does* suggest a fix for this. --- Mathlib/Algebra/Module/Torsion.lean | 4 ++-- Mathlib/LinearAlgebra/Dual.lean | 36 ++++++++++++++--------------- 2 files changed, 20 insertions(+), 20 deletions(-) diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index a993b226e188a..3121a844257b9 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -67,7 +67,7 @@ variable (R M : Type*) [Semiring R] [AddCommMonoid M] [Module R M] /-- The torsion ideal of `x`, containing all `a` such that `a • x = 0`. -/ @[simps!] def torsionOf (x : M) : Ideal R := - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1629 LinearMap.ker (LinearMap.toSpanSingleton R M x) @[simp] @@ -148,7 +148,7 @@ namespace Submodule `a • x = 0`. -/ @[simps!] def torsionBy (a : R) : Submodule R M := - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation on LinearMap.ker https://github.com/leanprover/lean4/issues/1629 LinearMap.ker (DistribMulAction.toLinearMap R M a) /-- The submodule containing all elements `x` of `M` such that `a • x = 0` for all `a` in `s`. -/ diff --git a/Mathlib/LinearAlgebra/Dual.lean b/Mathlib/LinearAlgebra/Dual.lean index 1ef8aede9c87b..a203905b0a8a2 100644 --- a/Mathlib/LinearAlgebra/Dual.lean +++ b/Mathlib/LinearAlgebra/Dual.lean @@ -339,11 +339,11 @@ theorem toDual_injective : Injective b.toDual := fun x y h ↦ b.ext_elem_iff.mp theorem toDual_inj (m : M) (a : b.toDual m = 0) : m = 0 := b.toDual_injective (by rwa [_root_.map_zero]) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem toDual_ker : LinearMap.ker b.toDual = ⊥ := ker_eq_bot'.mpr b.toDual_inj --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem toDual_range [Finite ι] : LinearMap.range b.toDual = ⊤ := by refine eq_top_iff'.2 fun f => ?_ let lin_comb : ι →₀ R := Finsupp.equivFunOnFinite.symm fun i => f (b i) @@ -447,7 +447,7 @@ theorem eval_ker {ι : Type*} (b : Basis ι R M) : simp_rw [LinearMap.ext_iff, Dual.eval_apply, zero_apply] at hm exact (Basis.forall_coord_eq_zero_iff _).mp fun i => hm (b.coord i) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem eval_range {ι : Type*} [Finite ι] (b : Basis ι R M) : LinearMap.range (Dual.eval R M) = ⊤ := by classical @@ -501,7 +501,7 @@ section variable (K) (V) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem eval_ker : LinearMap.ker (eval K V) = ⊥ := have ⟨s, hs⟩ := Module.projective_def'.mp ‹Projective K V› ker_eq_bot.mpr <| .of_comp (f := s.dualMap.dualMap) <| (ker_eq_bot.mp <| @@ -595,7 +595,7 @@ instance (priority := 900) IsReflexive.of_finite_of_free [Module.Finite R M] [Fr variable [IsReflexive R M] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem erange_coe : LinearMap.range (eval R M) = ⊤ := range_eq_top.mpr (bijective_dual_eval _ _).2 @@ -943,7 +943,7 @@ theorem dualRestrict_apply (W : Submodule R M) (φ : Module.Dual R M) (x : W) : that `φ w = 0` for all `w ∈ W`. -/ def dualAnnihilator {R : Type u} {M : Type v} [CommSemiring R] [AddCommMonoid M] [Module R M] (W : Submodule R M) : Submodule R <| Module.Dual R M := --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker LinearMap.ker W.dualRestrict @[simp] @@ -955,7 +955,7 @@ theorem mem_dualAnnihilator (φ : Module.Dual R M) : φ ∈ W.dualAnnihilator /-- That $\operatorname{ker}(\iota^* : V^* \to W^*) = \operatorname{ann}(W)$. This is the definition of the dual annihilator of the submodule $W$. -/ theorem dualRestrict_ker_eq_dualAnnihilator (W : Submodule R M) : - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker LinearMap.ker W.dualRestrict = W.dualAnnihilator := rfl @@ -1189,7 +1189,7 @@ theorem quotAnnihilatorEquiv_apply (W : Subspace K V) (φ : Module.Dual K V) : rfl /-- The natural isomorphism from the dual of a subspace `W` to `W.dualLift.range`. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range noncomputable def dualEquivDual (W : Subspace K V) : Module.Dual K W ≃ₗ[K] LinearMap.range W.dualLift := LinearEquiv.ofInjective _ dualLift_injective @@ -1228,7 +1228,7 @@ theorem dualAnnihilator_dualAnnihilator_eq (W : Subspace K V) : rwa [← OrderIso.symm_apply_eq] /-- The quotient by the dual is isomorphic to its dual annihilator. -/ --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range noncomputable def quotDualEquivAnnihilator (W : Subspace K V) : (Module.Dual K V ⧸ LinearMap.range W.dualLift) ≃ₗ[K] W.dualAnnihilator := LinearEquiv.quotEquivOfQuotEquiv <| LinearEquiv.trans W.quotAnnihilatorEquiv W.dualEquivDual @@ -1274,7 +1274,7 @@ variable {R : Type uR} [CommSemiring R] {M₁ : Type uM₁} {M₂ : Type uM₂} variable [AddCommMonoid M₁] [Module R M₁] [AddCommMonoid M₂] [Module R M₂] variable (f : M₁ →ₗ[R] M₂) --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker theorem ker_dualMap_eq_dualAnnihilator_range : LinearMap.ker f.dualMap = f.range.dualAnnihilator := by ext @@ -1282,7 +1282,7 @@ theorem ker_dualMap_eq_dualAnnihilator_range : ← SetLike.mem_coe, range_coe, Set.forall_mem_range] rfl --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_le_dualAnnihilator_ker : LinearMap.range f.dualMap ≤ f.ker.dualAnnihilator := by rintro _ ⟨ψ, rfl⟩ @@ -1338,7 +1338,7 @@ theorem dualPairing_apply {W : Submodule R M} (φ : Module.Dual R M) (x : W) : W.dualPairing (Quotient.mk φ) x = φ x := rfl --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range /-- That $\operatorname{im}(q^* : (V/W)^* \to V^*) = \operatorname{ann}(W)$. -/ theorem range_dualMap_mkQ_eq (W : Submodule R M) : LinearMap.range W.mkQ.dualMap = W.dualAnnihilator := by @@ -1360,7 +1360,7 @@ def dualQuotEquivDualAnnihilator (W : Submodule R M) : Module.Dual R (M ⧸ W) ≃ₗ[R] W.dualAnnihilator := LinearEquiv.ofLinear (W.mkQ.dualMap.codRestrict W.dualAnnihilator fun φ => --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.mem_range_self +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.mem_range_self W.range_dualMap_mkQ_eq ▸ LinearMap.mem_range_self W.mkQ.dualMap φ) W.dualCopairing (by ext; rfl) (by ext; rfl) @@ -1415,7 +1415,7 @@ namespace LinearMap open Submodule --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_eq_dualAnnihilator_ker_of_surjective (f : M →ₗ[R] M') (hf : Function.Surjective f) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := ((f.quotKerEquivOfSurjective hf).dualMap.range_comp _).trans f.ker.range_dualMap_mkQ_eq @@ -1429,7 +1429,7 @@ theorem range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective (f : M rw [← range_eq_top, range_rangeRestrict] have := range_dualMap_eq_dualAnnihilator_ker_of_surjective f.rangeRestrict rr_surj convert this using 1 - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 · calc _ = range ((range f).subtype.comp f.rangeRestrict).dualMap := by simp _ = _ := ?_ @@ -1525,7 +1525,7 @@ theorem dualMap_surjective_of_injective {f : V₁ →ₗ[K] V₂} (hf : Function have ⟨f', hf'⟩ := f.exists_leftInverse_of_injective (ker_eq_bot.mpr hf) ⟨φ.comp f', ext fun x ↦ congr(φ <| $hf' x)⟩ - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.range + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.range theorem range_dualMap_eq_dualAnnihilator_ker (f : V₁ →ₗ[K] V₂) : LinearMap.range f.dualMap = f.ker.dualAnnihilator := range_dualMap_eq_dualAnnihilator_ker_of_subtype_range_surjective f <| @@ -1576,7 +1576,7 @@ theorem dualAnnihilator_inf_eq (W W' : Subspace K V₁) : (W ⊓ W').dualAnnihilator = W.dualAnnihilator ⊔ W'.dualAnnihilator := by refine le_antisymm ?_ (sup_dualAnnihilator_le_inf W W') let F : V₁ →ₗ[K] (V₁ ⧸ W) × V₁ ⧸ W' := (Submodule.mkQ W).prod (Submodule.mkQ W') - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 LinearMap.ker + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 LinearMap.ker have : LinearMap.ker F = W ⊓ W' := by simp only [F, LinearMap.ker_prod, ker_mkQ] rw [← this, ← LinearMap.range_dualMap_eq_dualAnnihilator_ker] intro φ @@ -1632,7 +1632,7 @@ namespace LinearMap @[simp] theorem finrank_range_dualMap_eq_finrank_range (f : V₁ →ₗ[K] V₂) : - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1910 + -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation https://github.com/leanprover/lean4/issues/1629 finrank K (LinearMap.range f.dualMap) = finrank K (LinearMap.range f) := by rw [congr_arg dualMap (show f = (range f).subtype.comp f.rangeRestrict by rfl), ← dualMap_comp_dualMap, range_comp, From ff064a56e58e8ceb1c0a3febd60e701a1cf408fe Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Sat, 23 Nov 2024 06:30:20 +0000 Subject: [PATCH 057/250] chore: change `pp_nodot` adaptation note links to lean4 issue #1910 to refer to lean4 issue #6178 instead (#19385) These adaptation notes refer to the `pp_nodot` attribute not doing anything on certain declarations, due to their inability to be used with dot notation. However, now that leanprover/lean4#1910 has been fixed, the limiting factor has become the fact that the pretty printer doesn't work for these kinds of dot notation, which leanprover/lean4#6178 is an issue for. As such, the links to the blocking lean issue have been changed in this PR. --- Mathlib/Algebra/Symmetrized.lean | 2 +- Mathlib/Data/NNReal/Defs.lean | 2 +- Mathlib/Data/Real/Sqrt.lean | 2 +- 3 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Algebra/Symmetrized.lean b/Mathlib/Algebra/Symmetrized.lean index 00c398aaf91dd..45a69d0768566 100644 --- a/Mathlib/Algebra/Symmetrized.lean +++ b/Mathlib/Algebra/Symmetrized.lean @@ -48,7 +48,7 @@ def sym : α ≃ αˢʸᵐ := /-- The element of `α` represented by `x : αˢʸᵐ`. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] def unsym : αˢʸᵐ ≃ α := Equiv.refl _ diff --git a/Mathlib/Data/NNReal/Defs.lean b/Mathlib/Data/NNReal/Defs.lean index 210cfbe55782c..9c117c2ed6f81 100644 --- a/Mathlib/Data/NNReal/Defs.lean +++ b/Mathlib/Data/NNReal/Defs.lean @@ -917,7 +917,7 @@ namespace Real /-- The absolute value on `ℝ` as a map to `ℝ≥0`. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] def nnabs : ℝ →*₀ ℝ≥0 where toFun x := ⟨|x|, abs_nonneg x⟩ diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index b07597c19deda..b43bd46488f69 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -39,7 +39,7 @@ variable {x y : ℝ≥0} /-- Square root of a nonnegative real number. -/ -- Porting note (kmill): `pp_nodot` has no affect here --- unless RFC https://github.com/leanprover/lean4/issues/1910 leads to dot notation for CoeFun +-- unless RFC https://github.com/leanprover/lean4/issues/6178 leads to dot notation pp for CoeFun @[pp_nodot] noncomputable def sqrt : ℝ≥0 ≃o ℝ≥0 := OrderIso.symm <| powOrderIso 2 two_ne_zero From 4b6c6f8752e065dc0040b8165cfb7f76f1e87467 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 06:55:58 +0000 Subject: [PATCH 058/250] chore: cleanup of nolint simpNF (#19061) --- Mathlib/Algebra/Algebra/Unitization.lean | 3 +-- Mathlib/Algebra/BigOperators/Finsupp.lean | 14 +++++--------- Mathlib/Algebra/Group/WithOne/Basic.lean | 3 +-- .../Algebra/Homology/ShortComplex/Homology.lean | 2 -- .../Homology/ShortComplex/LeftHomology.lean | 1 - .../Homology/ShortComplex/RightHomology.lean | 1 - Mathlib/Algebra/Lie/IdealOperations.lean | 2 -- Mathlib/Algebra/Module/GradedModule.lean | 4 +--- Mathlib/Algebra/Module/Submodule/Pointwise.lean | 4 +--- Mathlib/Algebra/MvPolynomial/Basic.lean | 5 ++--- Mathlib/Algebra/Ring/Semiconj.lean | 8 ++------ Mathlib/Algebra/Star/Subsemiring.lean | 1 - Mathlib/RingTheory/DedekindDomain/Ideal.lean | 4 ---- 13 files changed, 13 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Unitization.lean b/Mathlib/Algebra/Algebra/Unitization.lean index 42a25ab336b19..89f706b546d03 100644 --- a/Mathlib/Algebra/Algebra/Unitization.lean +++ b/Mathlib/Algebra/Algebra/Unitization.lean @@ -743,8 +743,7 @@ def starLift : (A →⋆ₙₐ[R] C) ≃ (Unitization R A →⋆ₐ[R] C) := right_inv := fun φ => Unitization.algHom_ext'' <| by simp } --- Note (#6057) : tagging simpNF because linter complains -@[simp high, nolint simpNF] +@[simp high] theorem starLift_symm_apply_apply (φ : Unitization R A →⋆ₐ[R] C) (a : A) : Unitization.starLift.symm φ a = φ a := rfl diff --git a/Mathlib/Algebra/BigOperators/Finsupp.lean b/Mathlib/Algebra/BigOperators/Finsupp.lean index a02b2d8a48142..dffd8f94b48be 100644 --- a/Mathlib/Algebra/BigOperators/Finsupp.lean +++ b/Mathlib/Algebra/BigOperators/Finsupp.lean @@ -86,22 +86,17 @@ theorem prod_ite_eq [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq] -/- Porting note: simpnf linter, added aux lemma below -Left-hand side simplifies from - Finsupp.sum f fun x v => if a = x then v else 0 -to - if ↑f a = 0 then 0 else ↑f a --/ --- @[simp] theorem sum_ite_self_eq [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (a = x) v 0) = f a := by classical convert f.sum_ite_eq a fun _ => id simp [ite_eq_right_iff.2 Eq.symm] --- Porting note: Added this thm to replace the simp in the previous one. Need to add [DecidableEq N] +/-- +The left hand side of `sum_ite_self_eq` simplifies; this is the variant that is useful for `simp`. +-/ @[simp] -theorem sum_ite_self_eq_aux [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : +theorem if_mem_support [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (if a ∈ f.support then f a else 0) = f a := by simp only [mem_support_iff, ne_eq, ite_eq_left_iff, not_not] exact fun h ↦ h.symm @@ -113,6 +108,7 @@ theorem prod_ite_eq' [DecidableEq α] (f : α →₀ M) (a : α) (b : α → M dsimp [Finsupp.prod] rw [f.support.prod_ite_eq'] +/-- A restatement of `sum_ite_self_eq` with the equality test reversed. -/ theorem sum_ite_self_eq' [DecidableEq α] {N : Type*} [AddCommMonoid N] (f : α →₀ N) (a : α) : (f.sum fun x v => ite (x = a) v 0) = f a := by classical diff --git a/Mathlib/Algebra/Group/WithOne/Basic.lean b/Mathlib/Algebra/Group/WithOne/Basic.lean index 3dca5ba0dab77..b9d87ce1b04af 100644 --- a/Mathlib/Algebra/Group/WithOne/Basic.lean +++ b/Mathlib/Algebra/Group/WithOne/Basic.lean @@ -73,8 +73,7 @@ variable (f : α →ₙ* β) theorem lift_coe (x : α) : lift f x = f x := rfl --- Porting note (#11119): removed `simp` attribute to appease `simpNF` linter. -@[to_additive] +@[to_additive (attr := simp)] theorem lift_one : lift f 1 = 1 := rfl diff --git a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean index 25854245e340c..0c416a49f4df5 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/Homology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/Homology.lean @@ -70,8 +70,6 @@ structure HomologyMapData where namespace HomologyMapData -attribute [nolint simpNF] mk.injEq - variable {φ h₁ h₂} @[reassoc] diff --git a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean index 1022a52c7ab90..a8abcedd4972c 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/LeftHomology.lean @@ -266,7 +266,6 @@ structure LeftHomologyMapData where namespace LeftHomologyMapData attribute [reassoc (attr := simp)] commi commf' commπ -attribute [nolint simpNF] mk.injEq /-- The left homology map data associated to the zero morphism between two short complexes. -/ @[simps] diff --git a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean index 865ffa00b94fb..3592289533893 100644 --- a/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean +++ b/Mathlib/Algebra/Homology/ShortComplex/RightHomology.lean @@ -356,7 +356,6 @@ structure RightHomologyMapData where namespace RightHomologyMapData attribute [reassoc (attr := simp)] commp commg' commι -attribute [nolint simpNF] mk.injEq /-- The right homology map data associated to the zero morphism between two short complexes. -/ @[simps] diff --git a/Mathlib/Algebra/Lie/IdealOperations.lean b/Mathlib/Algebra/Lie/IdealOperations.lean index 75e9facacd878..459dce3dfdbf0 100644 --- a/Mathlib/Algebra/Lie/IdealOperations.lean +++ b/Mathlib/Algebra/Lie/IdealOperations.lean @@ -192,12 +192,10 @@ theorem sup_lie : ⁅I ⊔ J, N⁆ = ⁅I, N⁆ ⊔ ⁅J, N⁆ := by use ⁅((⟨x₂, hx₂⟩ : J) : L), (n : N)⁆; constructor; · apply lie_coe_mem_lie simp [← h, ← hx'] --- @[simp] -- Porting note: not in simpNF theorem lie_inf : ⁅I, N ⊓ N'⁆ ≤ ⁅I, N⁆ ⊓ ⁅I, N'⁆ := by rw [le_inf_iff]; constructor <;> apply mono_lie_right <;> [exact inf_le_left; exact inf_le_right] --- @[simp] -- Porting note: not in simpNF theorem inf_lie : ⁅I ⊓ J, N⁆ ≤ ⁅I, N⁆ ⊓ ⁅J, N⁆ := by rw [le_inf_iff]; constructor <;> apply mono_lie_left <;> [exact inf_le_left; exact inf_le_right] diff --git a/Mathlib/Algebra/Module/GradedModule.lean b/Mathlib/Algebra/Module/GradedModule.lean index 8adb7c6ada0a5..8797f2bbdd7b5 100644 --- a/Mathlib/Algebra/Module/GradedModule.lean +++ b/Mathlib/Algebra/Module/GradedModule.lean @@ -93,11 +93,9 @@ theorem smulAddMonoidHom_apply_of_of [DecidableEq ιA] [DecidableEq ιB] [GMonoi smulAddMonoidHom A M (DirectSum.of A i x) (of M j y) = of M (i +ᵥ j) (GSMul.smul x y) := by simp [smulAddMonoidHom] --- @[simp] -- Porting note: simpNF lint theorem of_smul_of [DecidableEq ιA] [DecidableEq ιB] [GMonoid A] [Gmodule A M] {i j} (x : A i) (y : M j) : - DirectSum.of A i x • of M j y = of M (i +ᵥ j) (GSMul.smul x y) := - smulAddMonoidHom_apply_of_of _ _ _ _ + DirectSum.of A i x • of M j y = of M (i +ᵥ j) (GSMul.smul x y) := by simp open AddMonoidHom diff --git a/Mathlib/Algebra/Module/Submodule/Pointwise.lean b/Mathlib/Algebra/Module/Submodule/Pointwise.lean index ec25efdc44f71..99f269138c452 100644 --- a/Mathlib/Algebra/Module/Submodule/Pointwise.lean +++ b/Mathlib/Algebra/Module/Submodule/Pointwise.lean @@ -161,9 +161,7 @@ instance pointwiseAddCommMonoid : AddCommMonoid (Submodule R M) where theorem add_eq_sup (p q : Submodule R M) : p + q = p ⊔ q := rfl --- dsimp loops when applying this lemma to its LHS, --- probably https://github.com/leanprover/lean4/pull/2867 -@[simp, nolint simpNF] +@[simp] theorem zero_eq_bot : (0 : Submodule R M) = ⊥ := rfl diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index dac73ae71fe7a..b73915560d966 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -844,9 +844,8 @@ theorem constantCoeff_X (i : σ) : constantCoeff (X i : MvPolynomial σ R) = 0 : simp [constantCoeff_eq] variable {R} -/- porting note: increased priority because otherwise `simp` time outs when trying to simplify -the left-hand side. `simpNF` linter indicated this and it was verified. -/ -@[simp 1001] + +@[simp] theorem constantCoeff_smul {R : Type*} [SMulZeroClass R S₁] (a : R) (f : MvPolynomial σ S₁) : constantCoeff (a • f) = a • constantCoeff f := rfl diff --git a/Mathlib/Algebra/Ring/Semiconj.lean b/Mathlib/Algebra/Ring/Semiconj.lean index 8b20e68a71a06..0ec9636ddebc6 100644 --- a/Mathlib/Algebra/Ring/Semiconj.lean +++ b/Mathlib/Algebra/Ring/Semiconj.lean @@ -61,13 +61,9 @@ section variable [MulOneClass R] [HasDistribNeg R] --- Porting note: `simpNF` told me to remove `simp` attribute -theorem neg_one_right (a : R) : SemiconjBy a (-1) (-1) := - (one_right a).neg_right +theorem neg_one_right (a : R) : SemiconjBy a (-1) (-1) := by simp --- Porting note: `simpNF` told me to remove `simp` attribute -theorem neg_one_left (x : R) : SemiconjBy (-1) x x := - (SemiconjBy.one_left x).neg_left +theorem neg_one_left (x : R) : SemiconjBy (-1) x x := by simp end diff --git a/Mathlib/Algebra/Star/Subsemiring.lean b/Mathlib/Algebra/Star/Subsemiring.lean index 8ef2bdd8fdd39..707ac11a068a8 100644 --- a/Mathlib/Algebra/Star/Subsemiring.lean +++ b/Mathlib/Algebra/Star/Subsemiring.lean @@ -53,7 +53,6 @@ instance starRing (s : StarSubsemiring R) : StarRing s := instance semiring (s : StarSubsemiring R) : NonAssocSemiring s := s.toSubsemiring.toNonAssocSemiring -@[simp, nolint simpNF] theorem mem_carrier {s : StarSubsemiring R} {x : R} : x ∈ s.carrier ↔ x ∈ s := Iff.rfl diff --git a/Mathlib/RingTheory/DedekindDomain/Ideal.lean b/Mathlib/RingTheory/DedekindDomain/Ideal.lean index 22365ec289413..2533459ec4771 100644 --- a/Mathlib/RingTheory/DedekindDomain/Ideal.lean +++ b/Mathlib/RingTheory/DedekindDomain/Ideal.lean @@ -129,7 +129,6 @@ open Submodule Submodule.IsPrincipal theorem spanSingleton_inv (x : K) : (spanSingleton R₁⁰ x)⁻¹ = spanSingleton _ x⁻¹ := one_div_spanSingleton x --- @[simp] -- Porting note: not in simpNF form theorem spanSingleton_div_spanSingleton (x y : K) : spanSingleton R₁⁰ x / spanSingleton R₁⁰ y = spanSingleton R₁⁰ (x / y) := by rw [div_spanSingleton, mul_comm, spanSingleton_mul_spanSingleton, div_eq_mul_inv] @@ -1058,7 +1057,6 @@ variable [IsDedekindDomain R] (f : R ⧸ I ≃+* A ⧸ J) /-- The bijection between ideals of `R` dividing `I` and the ideals of `A` dividing `J` induced by an isomorphism `f : R/I ≅ A/J`. -/ --- @[simps] -- Porting note: simpNF complains about the lemmas generated by simps def idealFactorsEquivOfQuotEquiv : { p : Ideal R | p ∣ I } ≃o { p : Ideal A | p ∣ J } := by have f_surj : Function.Surjective (f : R ⧸ I →+* A ⧸ J) := f.surjective have fsym_surj : Function.Surjective (f.symm : A ⧸ J →+* R ⧸ I) := f.symm.surjective @@ -1101,7 +1099,6 @@ theorem idealFactorsEquivOfQuotEquiv_mem_normalizedFactors_of_mem_normalizedFact /-- The bijection between the sets of normalized factors of I and J induced by a ring isomorphism `f : R/I ≅ A/J`. -/ --- @[simps apply] -- Porting note: simpNF complains about the lemmas generated by simps def normalizedFactorsEquivOfQuotEquiv (hI : I ≠ ⊥) (hJ : J ≠ ⊥) : { L : Ideal R | L ∈ normalizedFactors I } ≃ { M : Ideal A | M ∈ normalizedFactors J } where toFun j := @@ -1379,7 +1376,6 @@ variable [NormalizationMonoid R] /-- The bijection between the (normalized) prime factors of `r` and the (normalized) prime factors of `span {r}` -/ --- @[simps] -- Porting note: simpNF complains about the lemmas generated by simps noncomputable def normalizedFactorsEquivSpanNormalizedFactors {r : R} (hr : r ≠ 0) : { d : R | d ∈ normalizedFactors r } ≃ { I : Ideal R | I ∈ normalizedFactors (Ideal.span ({r} : Set R)) } := by From 409137130c4c0e7033eea0a7e369aa8607fd0973 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 23 Nov 2024 06:55:59 +0000 Subject: [PATCH 059/250] fix(Tactic/Linter): `upstreamableDecl` should treat `structure`s like `def`s (#19360) Reported on Zulip: https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/upstreamableDecl.20linter.20bug/near/483790949 We want to treat a `structure`/`inductive` declaration in the current file just like a `def` for the `upstreamableDecl` linter: potentially place a warning on the `structure` itself, but otherwise don't place warnings on downstream uses. There's a small complication where `inductive` declarations are reported to depend on their own constructors. I couldn't find a good way to determine whether any given constant is indeed a constructor declared in some given piece of syntax, so instead we allow depending on constructors just like we allow depending on theorems. Co-authored-by: Anne Baanen --- Mathlib/Tactic/Linter/UpstreamableDecl.lean | 20 ++++++++++++----- MathlibTest/MinImports.lean | 24 +++++++++++++++++++++ 2 files changed, 39 insertions(+), 5 deletions(-) diff --git a/Mathlib/Tactic/Linter/UpstreamableDecl.lean b/Mathlib/Tactic/Linter/UpstreamableDecl.lean index 86f55dc84e8bd..59beda016bc47 100644 --- a/Mathlib/Tactic/Linter/UpstreamableDecl.lean +++ b/Mathlib/Tactic/Linter/UpstreamableDecl.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Damiano Testa, Anne Baanen -/ import ImportGraph.Imports -import Mathlib.Lean.Expr.Basic import Mathlib.Tactic.MinImports /-! # The `upstreamableDecl` linter @@ -21,8 +20,12 @@ def Lean.Name.isLocal (env : Environment) (decl : Name) : Bool := open Mathlib.Command.MinImports -/-- Does the declaration with this name depend on `def`s in the current file? -/ -def Lean.Environment.localDefDependencies (env : Environment) (stx id : Syntax) : +/-- Does the declaration with this name depend on definitions in the current file? + +Here, "definition" means everything that is not a theorem, and so includes `def`, +`structure`, `inductive`, etc. +-/ +def Lean.Environment.localDefinitionDependencies (env : Environment) (stx id : Syntax) : CommandElabM Bool := do let declName : NameSet ← try NameSet.ofList <$> resolveGlobalConst id @@ -37,7 +40,14 @@ def Lean.Environment.localDefDependencies (env : Environment) (stx id : Syntax) let deps ← liftCoreM <| immediateDeps.transitivelyUsedConstants let constInfos := deps.toList.filterMap env.find? - let defs := constInfos.filter (·.isDef) + -- We allow depending on theorems and constructors. + -- We explicitly allow constructors since `inductive` declarations are reported to depend on their + -- own constructors, and we want inductives to behave the same as definitions, so place one + -- warning on the inductive itself but nothing on its downstream uses. + -- (There does not seem to be an easy way to determine, given `Syntax` and `ConstInfo`, + -- whether the `ConstInfo` is a constructor declared in this piece of `Syntax`.) + let defs := constInfos.filter (fun constInfo => !(constInfo.isTheorem || constInfo.isCtor)) + return defs.any fun constInfo => !(declName.contains constInfo.name) && constInfo.name.isLocal env namespace Mathlib.Linter @@ -72,7 +82,7 @@ def upstreamableDeclLinter : Linter where run := withSetOptionIn fun stx ↦ do let minImports := getIrredundantImports env (← getAllImports stx id) match minImports with | ⟨(RBNode.node _ .leaf upstream _ .leaf), _⟩ => do - if !(← env.localDefDependencies stx id) then + if !(← env.localDefinitionDependencies stx id) then let p : GoToModuleLinkProps := { modName := upstream } let widget : MessageData := .ofWidget (← liftCoreM <| Widget.WidgetInstance.ofHash diff --git a/MathlibTest/MinImports.lean b/MathlibTest/MinImports.lean index 43ee73fa011e9..87a0a011d07ad 100644 --- a/MathlibTest/MinImports.lean +++ b/MathlibTest/MinImports.lean @@ -165,4 +165,28 @@ def def_with_multiple_dependencies := let _ := Mathlib.Meta.NormNum.evalNatDvd false +/-! Structures and inductives should be treated just like definitions. -/ + +/-- + +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +structure ProposeToMoveThisStructure where + foo : ℕ + +/-- +warning: Consider moving this declaration to the module Mathlib.Data.Nat.Notation. +note: this linter can be disabled with `set_option linter.upstreamableDecl false` +-/ +#guard_msgs in +inductive ProposeToMoveThisInductive where +| foo : ℕ → ProposeToMoveThisInductive +| bar : ℕ → ProposeToMoveThisInductive → ProposeToMoveThisInductive + +-- This theorem depends on a local inductive, so should not be moved. +#guard_msgs in +theorem theorem_with_local_inductive : Nonempty ProposeToMoveThisInductive := ⟨.foo 0⟩ + end Linter.UpstreamableDecl From dde4596f8b160e9e0026507e6d34b40776bc2e65 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 07:36:13 +0000 Subject: [PATCH 060/250] chore: cleanup of many set_option deprecated.linter false (#19181) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Violeta Hernández --- Mathlib.lean | 2 +- Mathlib/Algebra/AddTorsor.lean | 10 ++++++++-- Mathlib/Algebra/BigOperators/Group/List.lean | 3 ++- .../Algebra/GroupWithZero/Units/Basic.lean | 4 ++-- Mathlib/Algebra/Ring/Parity.lean | 2 -- Mathlib/Data/FP/Basic.lean | 1 - Mathlib/Data/List/Basic.lean | 12 ++++++----- Mathlib/Data/List/Range.lean | 1 + Mathlib/Data/Multiset/Basic.lean | 7 +------ Mathlib/Data/Nat/Defs.lean | 1 - Mathlib/Data/Num/Lemmas.lean | 1 - Mathlib/Data/Num/Prime.lean | 1 - .../Basic.lean => Deprecated/LazyList.lean} | 0 .../SpecificGroups/Alternating.lean | 1 - .../AffineSpace/AffineEquiv.lean | 20 +++++++++++-------- .../CliffordAlgebra/Grading.lean | 13 +++++------- .../ExteriorAlgebra/Grading.lean | 10 ++-------- Mathlib/Logic/Denumerable.lean | 1 - Mathlib/Logic/Equiv/Nat.lean | 1 - .../Measure/MeasureSpaceDef.lean | 4 ++-- Mathlib/ModelTheory/Encoding.lean | 7 ------- Mathlib/Order/OmegaCompletePartialOrder.lean | 4 ++-- Mathlib/Order/RelClasses.lean | 2 -- Mathlib/SetTheory/Cardinal/Aleph.lean | 1 + Mathlib/SetTheory/Cardinal/Cofinality.lean | 1 - scripts/noshake.json | 1 + 26 files changed, 47 insertions(+), 64 deletions(-) rename Mathlib/{Data/LazyList/Basic.lean => Deprecated/LazyList.lean} (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 075b258277b1b..48be73827fdd6 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2507,7 +2507,6 @@ import Mathlib.Data.Int.Sqrt import Mathlib.Data.Int.Star import Mathlib.Data.Int.SuccPred import Mathlib.Data.Int.WithZero -import Mathlib.Data.LazyList.Basic import Mathlib.Data.List.AList import Mathlib.Data.List.Basic import Mathlib.Data.List.Chain @@ -2892,6 +2891,7 @@ import Mathlib.Deprecated.Combinator import Mathlib.Deprecated.Equiv import Mathlib.Deprecated.Group import Mathlib.Deprecated.HashMap +import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas diff --git a/Mathlib/Algebra/AddTorsor.lean b/Mathlib/Algebra/AddTorsor.lean index 3f6183df46993..5c7a2b005055b 100644 --- a/Mathlib/Algebra/AddTorsor.lean +++ b/Mathlib/Algebra/AddTorsor.lean @@ -404,13 +404,16 @@ theorem pointReflection_involutive (x : P) : Involutive (pointReflection x : P /-- `x` is the only fixed point of `pointReflection x`. This lemma requires `x + x = y + y ↔ x = y`. There is no typeclass to use here, so we add it as an explicit argument. -/ -theorem pointReflection_fixed_iff_of_injective_bit0 {x y : P} (h : Injective (2 • · : G → G)) : +theorem pointReflection_fixed_iff_of_injective_two_nsmul {x y : P} (h : Injective (2 • · : G → G)) : pointReflection x y = y ↔ y = x := by rw [pointReflection_apply, eq_comm, eq_vadd_iff_vsub_eq, ← neg_vsub_eq_vsub_rev, neg_eq_iff_add_eq_zero, ← two_nsmul, ← nsmul_zero 2, h.eq_iff, vsub_eq_zero_iff_eq, eq_comm] +@[deprecated (since := "2024-11-18")] alias pointReflection_fixed_iff_of_injective_bit0 := +pointReflection_fixed_iff_of_injective_two_nsmul + -- Porting note: need this to calm down CI -theorem injective_pointReflection_left_of_injective_bit0 {G P : Type*} [AddCommGroup G] +theorem injective_pointReflection_left_of_injective_two_nsmul {G P : Type*} [AddCommGroup G] [AddTorsor G P] (h : Injective (2 • · : G → G)) (y : P) : Injective fun x : P => pointReflection x y := fun x₁ x₂ (hy : pointReflection x₁ y = pointReflection x₂ y) => by @@ -418,6 +421,9 @@ theorem injective_pointReflection_left_of_injective_bit0 {G P : Type*} [AddCommG vsub_sub_vsub_cancel_right, ← neg_vsub_eq_vsub_rev, neg_eq_iff_add_eq_zero, ← two_nsmul, ← nsmul_zero 2, h.eq_iff, vsub_eq_zero_iff_eq] at hy +@[deprecated (since := "2024-11-18")] alias injective_pointReflection_left_of_injective_bit0 := +injective_pointReflection_left_of_injective_two_nsmul + end Equiv theorem AddTorsor.subsingleton_iff (G P : Type*) [AddGroup G] [AddTorsor G P] : diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 59dc2be68d10a..bf01489af4f70 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -657,7 +657,8 @@ theorem ranges_nodup {l s : List ℕ} (hs : s ∈ ranges l) : s.Nodup := /-- Any entry of any member of `l.ranges` is strictly smaller than `l.sum`. -/ lemma mem_mem_ranges_iff_lt_sum (l : List ℕ) {n : ℕ} : - (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by simp [mem_mem_ranges_iff_lt_natSum] + (∃ s ∈ l.ranges, n ∈ s) ↔ n < l.sum := by + rw [← mem_range, ← ranges_flatten, mem_flatten] @[simp] theorem length_flatMap (l : List α) (f : α → List β) : diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 4eed1a7034205..be6aebca0a776 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -422,9 +422,9 @@ instance (priority := 100) CommGroupWithZero.toDivisionCommMonoid : lemma div_mul_cancel_left₀ (ha : a ≠ 0) (b : G₀) : a / (a * b) = b⁻¹ := ha.isUnit.div_mul_cancel_left _ -set_option linter.deprecated false in @[deprecated div_mul_cancel_left₀ (since := "2024-03-22")] -lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := ha.isUnit.div_mul_right _ +lemma div_mul_right (b : G₀) (ha : a ≠ 0) : a / (a * b) = 1 / b := by + simp [div_mul_cancel_left₀ ha] lemma mul_div_cancel_left_of_imp (h : a = 0 → b = 0) : a * b / a = b := by rw [mul_comm, mul_div_cancel_of_imp h] diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index ff8df726aad0a..7b3959696657a 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -89,7 +89,6 @@ lemma Even.pow_of_ne_zero (ha : Even a) : ∀ {n : ℕ}, n ≠ 0 → Even (a ^ n /-- An element `a` of a semiring is odd if there exists `k` such `a = 2*k + 1`. -/ def Odd (a : α) : Prop := ∃ k, a = 2 * k + 1 -set_option linter.deprecated false in lemma odd_iff_exists_bit1 : Odd a ↔ ∃ b, a = 2 * b + 1 := exists_congr fun b ↦ by rw [two_mul] alias ⟨Odd.exists_bit1, _⟩ := odd_iff_exists_bit1 @@ -247,7 +246,6 @@ lemma mod_two_add_add_odd_mod_two (m : ℕ) {n : ℕ} (hn : Odd n) : m % 2 + (m lemma even_add' : Even (m + n) ↔ (Odd m ↔ Odd n) := by rw [even_add, ← not_odd_iff_even, ← not_odd_iff_even, not_iff_not] -set_option linter.deprecated false in @[simp] lemma not_even_bit1 (n : ℕ) : ¬Even (2 * n + 1) := by simp [parity_simps] lemma not_even_two_mul_add_one (n : ℕ) : ¬ Even (2 * n + 1) := diff --git a/Mathlib/Data/FP/Basic.lean b/Mathlib/Data/FP/Basic.lean index ebd1995796919..492cf4ffcd604 100644 --- a/Mathlib/Data/FP/Basic.lean +++ b/Mathlib/Data/FP/Basic.lean @@ -146,7 +146,6 @@ unsafe def nextUpPos (e m) (v : ValidFinite e m) : Float := Float.finite false e m' (by unfold ValidFinite at *; rw [ss]; exact v) else if h : e = emax then Float.inf false else Float.finite false e.succ (Nat.div2 m') lcProof -set_option linter.deprecated false in -- Porting note: remove this line when you dropped 'lcProof' set_option linter.unusedVariables false in @[nolint docBlame] diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index fac7238168886..4e3ce8533a890 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -737,7 +737,6 @@ end IndexOf /-! ### nth element -/ section deprecated -set_option linter.deprecated false @[simp] theorem getElem?_length (l : List α) : l[l.length]? = none := getElem?_len_le le_rfl @@ -754,19 +753,20 @@ theorem getElem_map_rev (f : α → β) {l} {n : Nat} {h : n < l.length} : /-- A version of `get_map` that can be used for rewriting. -/ @[deprecated getElem_map_rev (since := "2024-06-12")] theorem get_map_rev (f : α → β) {l n} : - f (get l n) = get (map f l) ⟨n.1, (l.length_map f).symm ▸ n.2⟩ := Eq.symm (get_map _) + f (get l n) = get (map f l) ⟨n.1, (l.length_map f).symm ▸ n.2⟩ := Eq.symm (getElem_map _) theorem get_length_sub_one {l : List α} (h : l.length - 1 < l.length) : l.get ⟨l.length - 1, h⟩ = l.getLast (by rintro rfl; exact Nat.lt_irrefl 0 h) := - (getLast_eq_get l _).symm + (getLast_eq_getElem l _).symm theorem take_one_drop_eq_of_lt_length {l : List α} {n : ℕ} (h : n < l.length) : (l.drop n).take 1 = [l.get ⟨n, h⟩] := by - rw [drop_eq_get_cons h, take, take] + rw [drop_eq_getElem_cons h, take, take] + simp theorem ext_get?' {l₁ l₂ : List α} (h' : ∀ n < max l₁.length l₂.length, l₁.get? n = l₂.get? n) : l₁ = l₂ := by - apply ext + apply ext_get? intro n rcases Nat.lt_or_ge n <| max l₁.length l₂.length with hn | hn · exact h' n hn @@ -831,6 +831,8 @@ theorem get_reverse (l : List α) (i : Nat) (h1 h2) : dsimp omega +set_option linter.deprecated false + theorem get_reverse' (l : List α) (n) (hn') : l.reverse.get n = l.get ⟨l.length - 1 - n, hn'⟩ := by rw [eq_comm] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 6a342b20f3573..7cefc522da8e2 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -149,6 +149,7 @@ lemma ranges_flatten' : ∀ l : List ℕ, l.ranges.flatten = range (Nat.sum l) set_option linter.deprecated false in /-- Any entry of any member of `l.ranges` is strictly smaller than `Nat.sum l`. See `List.mem_mem_ranges_iff_lt_sum` for the version about `List.sum`. -/ +@[deprecated "Use `List.mem_mem_ranges_iff_lt_sum`." (since := "2024-11-18")] lemma mem_mem_ranges_iff_lt_natSum (l : List ℕ) {n : ℕ} : (∃ s ∈ l.ranges, n ∈ s) ↔ n < Nat.sum l := by rw [← mem_range, ← ranges_flatten', mem_flatten] diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 49f764b115ba5..e364da53cefec 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -1720,15 +1720,10 @@ def filter (s : Multiset α) : Multiset α := theorem filter_zero : filter p 0 = 0 := rfl -#adaptation_note -/-- -Please re-enable the linter once we moved to `nightly-2024-06-22` or later. --/ -set_option linter.deprecated false in @[congr] theorem filter_congr {p q : α → Prop} [DecidablePred p] [DecidablePred q] {s : Multiset α} : (∀ x ∈ s, p x ↔ q x) → filter p s = filter q s := - Quot.inductionOn s fun _l h => congr_arg ofList <| filter_congr' <| by simpa using h + Quot.inductionOn s fun _l h => congr_arg ofList <| List.filter_congr <| by simpa using h @[simp] theorem filter_add (s t : Multiset α) : filter p (s + t) = filter p s + filter p t := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 2327e2c3bb37a..85ee5070c28b1 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1168,7 +1168,6 @@ lemma mul_add_mod' (a b c : ℕ) : (a * b + c) % b = c % b := by rw [Nat.mul_com lemma mul_add_mod_of_lt (h : c < b) : (a * b + c) % b = c := by rw [Nat.mul_add_mod', Nat.mod_eq_of_lt h] -set_option linter.deprecated false in @[simp] protected theorem not_two_dvd_bit1 (n : ℕ) : ¬2 ∣ 2 * n + 1 := by omega diff --git a/Mathlib/Data/Num/Lemmas.lean b/Mathlib/Data/Num/Lemmas.lean index 04f10c375c5b9..c7e98484c0e1c 100644 --- a/Mathlib/Data/Num/Lemmas.lean +++ b/Mathlib/Data/Num/Lemmas.lean @@ -308,7 +308,6 @@ theorem of_to_nat' : ∀ n : Num, Num.ofNat' (n : ℕ) = n | 0 => ofNat'_zero | pos p => p.of_to_nat' -set_option linter.deprecated false in lemma toNat_injective : Injective (castNum : Num → ℕ) := LeftInverse.injective of_to_nat' @[norm_cast] diff --git a/Mathlib/Data/Num/Prime.lean b/Mathlib/Data/Num/Prime.lean index 6b3a141d15139..c183f23216e12 100644 --- a/Mathlib/Data/Num/Prime.lean +++ b/Mathlib/Data/Num/Prime.lean @@ -37,7 +37,6 @@ def minFacAux (n : PosNum) : ℕ → PosNum → PosNum | fuel + 1, k => if n < k.bit1 * k.bit1 then n else if k.bit1 ∣ n then k.bit1 else minFacAux n fuel k.succ -set_option linter.deprecated false in theorem minFacAux_to_nat {fuel : ℕ} {n k : PosNum} (h : Nat.sqrt n < fuel + k.bit1) : (minFacAux n fuel k : ℕ) = Nat.minFacAux n k.bit1 := by induction' fuel with fuel ih generalizing k <;> rw [minFacAux, Nat.minFacAux] diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Deprecated/LazyList.lean similarity index 100% rename from Mathlib/Data/LazyList/Basic.lean rename to Mathlib/Deprecated/LazyList.lean diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 982ee49d373ef..8b0f82d484a0d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -78,7 +78,6 @@ theorem IsThreeCycle.mem_alternatingGroup {f : Perm α} (h : IsThreeCycle f) : f ∈ alternatingGroup α := Perm.mem_alternatingGroup.mpr h.sign -set_option linter.deprecated false in theorem finRotate_bit1_mem_alternatingGroup {n : ℕ} : finRotate (2 * n + 1) ∈ alternatingGroup (Fin (2 * n + 1)) := by rw [mem_alternatingGroup, sign_finRotate, pow_mul, pow_two, Int.units_mul_self, one_pow] diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean index 63e23ebbb7c77..75e11673b7d74 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineEquiv.lean @@ -486,22 +486,26 @@ theorem pointReflection_self (x : P₁) : pointReflection k x x = x := theorem pointReflection_involutive (x : P₁) : Involutive (pointReflection k x : P₁ → P₁) := Equiv.pointReflection_involutive x -set_option linter.deprecated false in /-- `x` is the only fixed point of `pointReflection x`. This lemma requires `x + x = y + y ↔ x = y`. There is no typeclass to use here, so we add it as an explicit argument. -/ -theorem pointReflection_fixed_iff_of_injective_bit0 {x y : P₁} (h : Injective (2 • · : V₁ → V₁)) : - pointReflection k x y = y ↔ y = x := - Equiv.pointReflection_fixed_iff_of_injective_bit0 h +theorem pointReflection_fixed_iff_of_injective_two_nsmul {x y : P₁} + (h : Injective (2 • · : V₁ → V₁)) : pointReflection k x y = y ↔ y = x := + Equiv.pointReflection_fixed_iff_of_injective_two_nsmul h + +@[deprecated (since := "2024-11-18")] alias pointReflection_fixed_iff_of_injective_bit0 := +pointReflection_fixed_iff_of_injective_two_nsmul -set_option linter.deprecated false in -theorem injective_pointReflection_left_of_injective_bit0 +theorem injective_pointReflection_left_of_injective_two_nsmul (h : Injective (2 • · : V₁ → V₁)) (y : P₁) : Injective fun x : P₁ => pointReflection k x y := - Equiv.injective_pointReflection_left_of_injective_bit0 h y + Equiv.injective_pointReflection_left_of_injective_two_nsmul h y + +@[deprecated (since := "2024-11-18")] alias injective_pointReflection_left_of_injective_bit0 := +injective_pointReflection_left_of_injective_two_nsmul theorem injective_pointReflection_left_of_module [Invertible (2 : k)] : ∀ y, Injective fun x : P₁ => pointReflection k x y := - injective_pointReflection_left_of_injective_bit0 k fun x y h => by + injective_pointReflection_left_of_injective_two_nsmul k fun x y h => by dsimp at h rwa [two_nsmul, two_nsmul, ← two_smul k x, ← two_smul k y, (isUnit_of_invertible (2 : k)).smul_left_cancel] at h diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean index 39224b85617fe..d88ad52680183 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/Grading.lean @@ -75,7 +75,6 @@ nonrec theorem GradedAlgebra.ι_sq_scalar (m : M) : rw [GradedAlgebra.ι_apply Q, DirectSum.of_mul_of, DirectSum.algebraMap_apply] exact DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext rfl <| ι_sq_scalar _ _) -set_option linter.deprecated false in theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : -- Porting note: added a second `by apply` lift Q ⟨by apply GradedAlgebra.ι Q, by apply GradedAlgebra.ι_sq_scalar Q⟩ x' = @@ -85,29 +84,27 @@ theorem GradedAlgebra.lift_ι_eq (i' : ZMod 2) (x' : evenOdd Q i') : induction hx' using Submodule.iSup_induction' with | mem i x hx => obtain ⟨i, rfl⟩ := i - -- Porting note: `dsimp only [Subtype.coe_mk] at hx` doesn't work, use `change` instead - change x ∈ LinearMap.range (ι Q) ^ i at hx + dsimp only [Subtype.coe_mk] at hx induction hx using Submodule.pow_induction_on_left' with | algebraMap r => rw [AlgHom.commutes, DirectSum.algebraMap_apply]; rfl | add x y i hx hy ihx ihy => - -- Note: in https://github.com/leanprover-community/mathlib4/pull/8386 `map_add` had to be specialized to avoid a timeout - -- (the definition was already very slow) - rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add] + rw [map_add, ihx, ihy, ← AddMonoidHom.map_add] rfl | mem_mul m hm i x hx ih => obtain ⟨_, rfl⟩ := hm - rw [AlgHom.map_mul, ih, lift_ι_apply, GradedAlgebra.ι_apply Q, DirectSum.of_mul_of] + rw [map_mul, ih, lift_ι_apply, GradedAlgebra.ι_apply Q, DirectSum.of_mul_of] refine DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext ?_ ?_) <;> dsimp only [GradedMonoid.mk, Subtype.coe_mk] · rw [Nat.succ_eq_add_one, add_comm, Nat.cast_add, Nat.cast_one] rfl | zero => + set_option linter.deprecated false in rw [AlgHom.map_zero] apply Eq.symm apply DFinsupp.single_eq_zero.mpr; rfl | add x y hx hy ihx ihy => - rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl + rw [map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl /-- The clifford algebra is graded by the even and odd parts. -/ instance gradedAlgebra : GradedAlgebra (evenOdd Q) := diff --git a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean index 81a29f1543817..d39a95c8a5d29 100644 --- a/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean +++ b/Mathlib/LinearAlgebra/ExteriorAlgebra/Grading.lean @@ -54,22 +54,16 @@ def GradedAlgebra.liftι : ExteriorAlgebra R M →ₐ[R] ⨁ i : ℕ, ⋀[R]^i M := lift R ⟨by apply GradedAlgebra.ι R M, GradedAlgebra.ι_sq_zero R M⟩ -set_option linter.deprecated false in theorem GradedAlgebra.liftι_eq (i : ℕ) (x : ⋀[R]^i M) : GradedAlgebra.liftι R M x = DirectSum.of (fun i => ⋀[R]^i M) i x := by cases' x with x hx dsimp only [Subtype.coe_mk, DirectSum.lof_eq_of] - -- Porting note: original statement was - -- refine Submodule.pow_induction_on_left' _ (fun r => ?_) (fun x y i hx hy ihx ihy => ?_) - -- (fun m hm i x hx ih => ?_) hx - -- but it created invalid goals induction hx using Submodule.pow_induction_on_left' with | algebraMap => simp_rw [AlgHom.commutes, DirectSum.algebraMap_apply]; rfl - -- FIXME: specialized `map_add` to avoid a (whole-declaration) timeout - | add _ _ _ _ _ ihx ihy => simp_rw [AlgHom.map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl + | add _ _ _ _ _ ihx ihy => simp_rw [map_add, ihx, ihy, ← AddMonoidHom.map_add]; rfl | mem_mul _ hm _ _ _ ih => obtain ⟨_, rfl⟩ := hm - simp_rw [AlgHom.map_mul, ih, GradedAlgebra.liftι, lift_ι_apply, GradedAlgebra.ι_apply R M, + simp_rw [map_mul, ih, GradedAlgebra.liftι, lift_ι_apply, GradedAlgebra.ι_apply R M, DirectSum.of_mul_of] exact DirectSum.of_eq_of_gradedMonoid_eq (Sigma.subtype_ext (add_comm _ _) rfl) diff --git a/Mathlib/Logic/Denumerable.lean b/Mathlib/Logic/Denumerable.lean index f96dccd8f839c..d6160e1858ab6 100644 --- a/Mathlib/Logic/Denumerable.lean +++ b/Mathlib/Logic/Denumerable.lean @@ -121,7 +121,6 @@ instance option : Denumerable (Option α) := · rw [decode_option_succ, decode_eq_ofNat, Option.map_some', Option.mem_def] rw [encode_some, encode_ofNat]⟩ -set_option linter.deprecated false in /-- If `α` and `β` are denumerable, then so is their sum. -/ instance sum : Denumerable (α ⊕ β) := ⟨fun n => by diff --git a/Mathlib/Logic/Equiv/Nat.lean b/Mathlib/Logic/Equiv/Nat.lean index 0bc17024d2609..69466efcdac2d 100644 --- a/Mathlib/Logic/Equiv/Nat.lean +++ b/Mathlib/Logic/Equiv/Nat.lean @@ -36,7 +36,6 @@ def boolProdNatEquivNat : Bool × ℕ ≃ ℕ where def natSumNatEquivNat : ℕ ⊕ ℕ ≃ ℕ := (boolProdEquivSum ℕ).symm.trans boolProdNatEquivNat -set_option linter.deprecated false in @[simp] theorem natSumNatEquivNat_apply : ⇑natSumNatEquivNat = Sum.elim (2 * ·) (2 * · + 1) := by ext (x | x) <;> rfl diff --git a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean index 8eb64f43a20e2..11481cc61993f 100644 --- a/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean +++ b/Mathlib/MeasureTheory/Measure/MeasureSpaceDef.lean @@ -82,9 +82,9 @@ instance Measure.instFunLike [MeasurableSpace α] : FunLike (Measure α) (Set α coe μ := μ.toOuterMeasure coe_injective' | ⟨_, _, _⟩, ⟨_, _, _⟩, h => toOuterMeasure_injective <| DFunLike.coe_injective h -set_option linter.deprecated false in -- Not immediately obvious how to use `measure_empty` here. + instance Measure.instOuterMeasureClass [MeasurableSpace α] : OuterMeasureClass (Measure α) α where - measure_empty m := m.empty' + measure_empty m := measure_empty (μ := m.toOuterMeasure) measure_iUnion_nat_le m := m.iUnion_nat measure_mono m := m.mono diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 5395d5324a5a6..180cb86b870ed 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -181,13 +181,6 @@ or returns `default` if not possible. -/ def sigmaImp : (Σn, L.BoundedFormula α n) → (Σn, L.BoundedFormula α n) → Σn, L.BoundedFormula α n | ⟨m, φ⟩, ⟨n, ψ⟩ => if h : m = n then ⟨m, φ.imp (Eq.mp (by rw [h]) ψ)⟩ else default -#adaptation_note -/-- -`List.drop_sizeOf_le` is deprecated. -See https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/Why.20is.20.60Mathlib.2EModelTheory.2EEncoding.60.20using.20.60SizeOf.2EsizeOf.60.3F -for discussion about adapting this code. --/ -set_option linter.deprecated false in /-- Decodes a list of symbols as a list of formulas. -/ @[simp] lemma sigmaImp_apply {n} {φ ψ : L.BoundedFormula α n} : diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 309394381dfc2..4c97970762171 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -802,11 +802,11 @@ theorem seq_continuous' {β γ : Type v} (f : α → Part (β → γ)) (g : α intro apply map_continuous' _ _ hg +set_option linter.deprecated true + theorem continuous (F : α →𝒄 β) (C : Chain α) : F (ωSup C) = ωSup (C.map F) := F.ωScottContinuous.map_ωSup _ -set_option linter.deprecated true - /-- Construct a continuous function from a bare function, a continuous function, and a proof that they are equal. -/ -- Porting note: removed `@[reducible]` diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 6fb587d189e2c..c99abc227fd09 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -17,8 +17,6 @@ extend `LE` and/or `LT` while these classes take a relation as an explicit argum -/ -set_option linter.deprecated false - universe u v variable {α : Type u} {β : Type v} {r : α → α → Prop} {s : β → β → Prop} diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 3970c6fc6cb57..dafc9ee33dced 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -616,6 +616,7 @@ alias aleph'_omega := aleph'_omega0 def aleph'Equiv : Ordinal ≃ Cardinal := ⟨aleph', alephIdx, alephIdx_aleph', aleph'_alephIdx⟩ +@[deprecated aleph_eq_preAleph (since := "2024-10-22")] theorem aleph_eq_aleph' (o : Ordinal) : ℵ_ o = preAleph (ω + o) := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 4024caa0c1f7a..87b3654599332 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -287,7 +287,6 @@ theorem lsub_lt_ord {ι} {f : ι → Ordinal} {c : Ordinal} (hι : #ι < c.cof) (∀ i, f i < c) → lsub.{u, u} f < c := lsub_lt_ord_lift (by rwa [(#ι).lift_id]) -set_option linter.deprecated false in theorem cof_iSup_le_lift {ι} {f : ι → Ordinal} (H : ∀ i, f i < iSup f) : cof (iSup f) ≤ Cardinal.lift.{v, u} #ι := by rw [← Ordinal.sup] at * diff --git a/scripts/noshake.json b/scripts/noshake.json index 736e94310eaf9..86cfab061a740 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -364,6 +364,7 @@ "Mathlib.Util.Superscript"], "Mathlib.Deprecated.NatLemmas": ["Batteries.Data.Nat.Lemmas", "Batteries.WF"], "Mathlib.Deprecated.MinMax": ["Mathlib.Order.MinMax"], + "Mathlib.Deprecated.LazyList": ["Mathlib.Lean.Thunk"], "Mathlib.Deprecated.ByteArray": ["Batteries.Data.ByteSubarray"], "Mathlib.Data.Vector.Basic": ["Mathlib.Control.Applicative"], "Mathlib.Data.Set.Image": From 84d03da261895eac598b0ac33120abfe8ee711f0 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 08:16:21 +0000 Subject: [PATCH 061/250] feat(LinearAlgebra/QuadraticForm/Basic): weaken invertibility hypothesis on 2 (#14986) The aim of this is to eventually enable base-change of quadratic maps from a (not necessarily (torsion-)free) abelian group to an abelian group on which 2 is invertible (e.g., the real numbers). The current set-up requires 2 to be invertible in the base ring (which would be the integers in the situation above). This PR weakens the assumption for the definition of `QuadraticMap.associatedHom` and `QuadraticMap.associated` from `[Invertible (2 : R)]` to `[Invertible (2 : Module.End R N)]`, where `N` is the target module of he quadratic map. See [this discussion](https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Quadratic.20forms.20with.20values.20in.20a.20larger.20ring/near/450582700) on Zulip. On the way, we fix some mistakes in docstrings. --- .../LinearAlgebra/QuadraticForm/Basic.lean | 163 ++++++++++++------ Mathlib/LinearAlgebra/QuadraticForm/Dual.lean | 7 +- 2 files changed, 119 insertions(+), 51 deletions(-) diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean index bec420998fa64..9778e86170c13 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Basic.lean @@ -526,8 +526,8 @@ variable [CommSemiring R] [CommSemiring S] [AddCommMonoid M] [Module R M] [AddCo [Module R N] [Module S M] [Module S N] [Algebra S R] variable [IsScalarTower S R M] [IsScalarTower S R N] -/-- If `B : M → N → Pₗ` is `R`-`S` bilinear and `R'` and `S'` are compatible scalar multiplications, -then the restriction of scalars is a `R'`-`S'` bilinear map. -/ +/-- If `Q : M → N` is a quadratic map of `R`-modules and `R` is an `S`-algebra, +then the restriction of scalars is a quadratic map of `S`-modules. -/ @[simps!] def restrictScalars (Q : QuadraticMap R M N) : QuadraticMap S M N where toFun x := Q x @@ -597,12 +597,13 @@ theorem _root_.LinearEquiv.congrQuadraticMap_symm (e : N ≃ₗ[R] P) : (LinearEquiv.congrQuadraticMap e (M := M)).symm = e.symm.congrQuadraticMap := rfl end Comp + section NonUnitalNonAssocSemiring variable [CommSemiring R] [NonUnitalNonAssocSemiring A] [AddCommMonoid M] [Module R M] variable [Module R A] [SMulCommClass R A A] [IsScalarTower R A A] -/-- The product of linear forms is a quadratic form. -/ +/-- The product of linear maps into an `R`-algebra is a quadratic map. -/ def linMulLin (f g : M →ₗ[R] A) : QuadraticMap R M A where toFun := f * g toFun_smul a x := by @@ -635,12 +636,12 @@ theorem linMulLin_comp (f g : M →ₗ[R] A) (h : N' →ₗ[R] M) : variable {n : Type*} -/-- `sq` is the quadratic form mapping the vector `x : A` to `x * x` -/ +/-- `sq` is the quadratic map sending the vector `x : A` to `x * x` -/ @[simps!] def sq : QuadraticMap R A A := linMulLin LinearMap.id LinearMap.id -/-- `proj i j` is the quadratic form mapping the vector `x : n → R` to `x i * x j` -/ +/-- `proj i j` is the quadratic map sending the vector `x : n → R` to `x i * x j` -/ def proj (i j : n) : QuadraticMap R (n → A) A := linMulLin (@LinearMap.proj _ _ _ (fun _ => A) _ _ i) (@LinearMap.proj _ _ _ (fun _ => A) _ _ j) @@ -655,10 +656,16 @@ end QuadraticMap /-! ### Associated bilinear maps -Over a commutative ring with an inverse of 2, the theory of quadratic maps is -basically identical to that of symmetric bilinear maps. The map from quadratic -maps to bilinear maps giving this identification is called the `QuadraticMap.associated` -quadratic map. +If multiplication by 2 is invertible on the target module `N` of +`QuadraticMap R M N`, then there is a linear bijection `QuadraticMap.associated` +between quadratic maps `Q` over `R` from `M` to `N` and symmetric bilinear maps +`B : M →ₗ[R] M →ₗ[R] → N` such that `BilinMap.toQuadraticMap B = Q` +(see `QuadraticMap.associated_rightInverse`). The associated bilinear map is half +`Q.polarBilin` (see `QuadraticMap.two_nsmul_associated`); this is where the invertibility condition +comes from. We spell the condition as `[Invertible (2 : Module.End R N)]`. + +Note that this makes the bijection available in more cases than the simpler condition +`Invertible (2 : R)`, e.g., when `R = ℤ` and `N = ℝ`. -/ namespace LinearMap @@ -711,14 +718,14 @@ section variable (S R M) -/-- `LinearMap.BilinForm.toQuadraticMap` as an additive homomorphism -/ +/-- `LinearMap.BilinMap.toQuadraticMap` as an additive homomorphism -/ @[simps] def toQuadraticMapAddMonoidHom : (BilinMap R M N) →+ QuadraticMap R M N where toFun := toQuadraticMap map_zero' := toQuadraticMap_zero _ _ map_add' := toQuadraticMap_add -/-- `LinearMap.BilinForm.toQuadraticMap` as a linear map -/ +/-- `LinearMap.BilinMap.toQuadraticMap` as a linear map -/ @[simps!] def toQuadraticMapLinearMap [Semiring S] [Module S N] [SMulCommClass S R N] [SMulCommClass R S N] : (BilinMap R M N) →ₗ[S] QuadraticMap R M N where @@ -818,43 +825,87 @@ namespace QuadraticMap open LinearMap (BilinMap) +section + +variable [Semiring R] [AddCommMonoid M] [Module R M] + +instance : SMulCommClass R (Submonoid.center R) M where + smul_comm r r' m := by + simp_rw [Submonoid.smul_def, smul_smul, (Set.mem_center_iff.1 r'.prop).1] + +/-- If `2` is invertible in `R`, then it is also invertible in `End R M`. -/ +instance [Invertible (2 : R)] : Invertible (2 : Module.End R M) where + invOf := (⟨⅟2, Set.invOf_mem_center (Set.ofNat_mem_center _ _)⟩ : Submonoid.center R) • + (1 : Module.End R M) + invOf_mul_self := by + ext m + dsimp [Submonoid.smul_def] + rw [← ofNat_smul_eq_nsmul R, invOf_smul_smul (2 : R) m] + mul_invOf_self := by + ext m + dsimp [Submonoid.smul_def] + rw [← ofNat_smul_eq_nsmul R, smul_invOf_smul (2 : R) m] + +/-- If `2` is invertible in `R`, then applying the inverse of `2` in `End R M` to an element +of `M` is the same as multiplying by the inverse of `2` in `R`. -/ +@[simp] +lemma half_moduleEnd_apply_eq_half_smul [Invertible (2 : R)] (x : M) : + ⅟ (2 : Module.End R M) x = ⅟ (2 : R) • x := + rfl + +end + section AssociatedHom -variable [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -variable (S) [CommSemiring S] [Algebra S R] -variable [Module S N] [IsScalarTower S R N] -variable [Invertible (2 : R)] {B₁ : BilinMap R M R} +variable [CommRing R] [AddCommGroup M] [Module R M] +variable [AddCommGroup N] [Module R N] +variable (S) [CommSemiring S] [Algebra S R] [Module S N] [IsScalarTower S R N] + +-- the requirement that multiplication by `2` is invertible on the target module `N` +variable [Invertible (2 : Module.End R N)] /-- `associatedHom` is the map that sends a quadratic map on a module `M` over `R` to its associated symmetric bilinear map. As provided here, this has the structure of an `S`-linear map -where `S` is a commutative subring of `R`. +where `S` is a commutative ring and `R` is an `S`-algebra. Over a commutative ring, use `QuadraticMap.associated`, which gives an `R`-linear map. Over a general ring with no nontrivial distinguished commutative subring, use `QuadraticMap.associated'`, which gives an additive homomorphism (or more precisely a `ℤ`-linear map.) -/ -def associatedHom : QuadraticMap R M N →ₗ[S] (BilinMap R M N) := - -- TODO: this `center` stuff is vertigial from an incorrect non-commutative version, but we leave - -- it behind to make a future refactor to a *correct* non-commutative version easier in future. - (⟨⅟2, Set.invOf_mem_center (Set.ofNat_mem_center _ _)⟩ : Submonoid.center R) • - { toFun := polarBilin - map_add' := fun _x _y => LinearMap.ext₂ <| polar_add _ _ - map_smul' := fun _c _x => LinearMap.ext₂ <| polar_smul _ _ } +def associatedHom : QuadraticMap R M N →ₗ[S] (BilinMap R M N) where + toFun Q := ⅟ (2 : Module.End R N) • polarBilin Q + map_add' _ _ := + LinearMap.ext₂ fun _ _ ↦ by + simp only [LinearMap.smul_apply, polarBilin_apply_apply, coeFn_add, polar_add, + LinearMap.smul_def, LinearMap.map_add, LinearMap.add_apply] + map_smul' _ _ := + LinearMap.ext₂ fun _ _ ↦ by + simp only [LinearMap.smul_apply, polarBilin_apply_apply, coeFn_smul, polar_smul, + LinearMap.smul_def, LinearMap.map_smul_of_tower, RingHom.id_apply] variable (Q : QuadraticMap R M N) @[simp] -theorem associated_apply (x y : M) : associatedHom S Q x y = ⅟ (2 : R) • (Q (x + y) - Q x - Q y) := +theorem associated_apply (x y : M) : + associatedHom S Q x y = ⅟ (2 : Module.End R N) • (Q (x + y) - Q x - Q y) := rfl +/-- Twice the associated bilinear map of `Q` is the same as the polar of `Q`. -/ @[simp] theorem two_nsmul_associated : 2 • associatedHom S Q = Q.polarBilin := by ext dsimp - rw [← smul_assoc, two_nsmul, invOf_two_add_invOf_two, one_smul, polar] + rw [← LinearMap.smul_apply, nsmul_eq_mul, Nat.cast_ofNat, mul_invOf_self', LinearMap.one_apply, + polar] -theorem associated_isSymm (Q : QuadraticMap R M R) : (associatedHom S Q).IsSymm := fun x y => by - simp only [associated_apply, sub_eq_add_neg, add_assoc, map_mul, RingHom.id_apply, map_add, - _root_.map_neg, add_comm, add_left_comm] +theorem associated_isSymm (Q : QuadraticForm R M) [Invertible (2 : R)] : + (associatedHom S Q).IsSymm := fun x y ↦ by + simp only [associated_apply, sub_eq_add_neg, add_assoc, RingHom.id_apply, add_comm, add_left_comm] +/-- A version of `QuadraticMap.associated_isSymm` for general targets +(using `flip` because `IsSymm` does not apply here). -/ +lemma associated_flip : (associatedHom S Q).flip = associatedHom S Q := by + ext + simp only [LinearMap.flip_apply, associated_apply, add_comm, sub_eq_add_neg, add_left_comm, + add_assoc] @[simp] theorem associated_comp {N' : Type*} [AddCommGroup N'] [Module R N'] (f : N' →ₗ[R] M) : @@ -862,23 +913,29 @@ theorem associated_comp {N' : Type*} [AddCommGroup N'] [Module R N'] (f : N' → ext simp only [associated_apply, comp_apply, map_add, LinearMap.compl₁₂_apply] -theorem associated_toQuadraticMap (B : BilinMap R M R) (x y : M) : - associatedHom S B.toQuadraticMap x y = ⅟ (2 : R) • (B x y + B y x) := by - simp only [associated_apply, LinearMap.BilinMap.toQuadraticMap_apply, map_add, - LinearMap.add_apply, smul_eq_mul] +theorem associated_toQuadraticMap (B : BilinMap R M N) (x y : M) : + associatedHom S B.toQuadraticMap x y = ⅟ (2 : Module.End R N) • (B x y + B y x) := by + simp only [associated_apply, BilinMap.toQuadraticMap_apply, map_add, LinearMap.add_apply, + LinearMap.smul_def, _root_.map_sub] abel_nf -theorem associated_left_inverse (h : B₁.IsSymm) : associatedHom S B₁.toQuadraticMap = B₁ := - LinearMap.ext₂ fun x y => by - rw [associated_toQuadraticMap, ← h.eq x y, RingHom.id_apply] - match_scalars - linear_combination invOf_mul_self' (2:R) +theorem associated_left_inverse [Invertible (2 : R)] {B₁ : BilinMap R M R} (h : B₁.IsSymm) : + associatedHom S B₁.toQuadraticMap = B₁ := + LinearMap.ext₂ fun x y ↦ by + rw [associated_toQuadraticMap, ← h.eq x y, RingHom.id_apply, ← two_mul, ← smul_eq_mul, + invOf_smul_eq_iff, two_smul, two_smul] + +/-- A version of `QuadraticMap.associated_left_inverse` for general targets. -/ +lemma associated_left_inverse' {B₁ : BilinMap R M N} (hB₁ : B₁.flip = B₁) : + associatedHom S B₁.toQuadraticMap = B₁ := by + ext _ y + rw [associated_toQuadraticMap, ← LinearMap.flip_apply _ y, hB₁, invOf_smul_eq_iff, two_smul] -- Porting note: moved from below to golf the next theorem theorem associated_eq_self_apply (x : M) : associatedHom S Q x x = Q x := by - rw [associated_apply, map_add_self] - match_scalars - linear_combination invOf_mul_self' (2:R) + rw [associated_apply, map_add_self, ← three_add_one_eq_four, ← two_add_one_eq_three, add_smul, + add_smul, one_smul, add_sub_cancel_right, add_sub_cancel_right, two_smul, ← two_smul R, + invOf_smul_eq_iff, two_smul, two_smul] theorem toQuadraticMap_associated : (associatedHom S Q).toQuadraticMap = Q := QuadraticMap.ext <| associated_eq_self_apply S Q @@ -887,7 +944,7 @@ theorem toQuadraticMap_associated : (associatedHom S Q).toQuadraticMap = Q := -- with historical naming in this file. theorem associated_rightInverse : Function.RightInverse (associatedHom S) (BilinMap.toQuadraticMap : _ → QuadraticMap R M N) := - fun Q => toQuadraticMap_associated S Q + toQuadraticMap_associated S /-- `associated'` is the `ℤ`-linear map that sends a quadratic form on a module `M` over `R` to its associated symmetric bilinear form. -/ @@ -895,10 +952,15 @@ abbrev associated' : QuadraticMap R M N →ₗ[ℤ] BilinMap R M N := associatedHom ℤ /-- Symmetric bilinear forms can be lifted to quadratic forms -/ -instance canLift : +instance canLift [Invertible (2 : R)] : CanLift (BilinMap R M R) (QuadraticForm R M) (associatedHom ℕ) LinearMap.IsSymm where prf B hB := ⟨B.toQuadraticMap, associated_left_inverse _ hB⟩ +/-- Symmetric bilinear maps can be lifted to quadratic maps -/ +instance canLift' : + CanLift (BilinMap R M N) (QuadraticMap R M N) (associatedHom ℕ) fun B ↦ B.flip = B where + prf B hB := ⟨B.toQuadraticMap, associated_left_inverse' _ hB⟩ + /-- There exists a non-null vector with respect to any quadratic form `Q` whose associated bilinear form is non-zero, i.e. there exists `x` such that `Q x ≠ 0`. -/ theorem exists_quadraticMap_ne_zero {Q : QuadraticMap R M N} @@ -919,10 +981,11 @@ section Associated variable [CommSemiring S] [CommRing R] [AddCommGroup M] [Algebra S R] [Module R M] variable [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower S R N] -variable [Invertible (2 : R)] +variable [Invertible (2 : Module.End R N)] -- Note: When possible, rather than writing lemmas about `associated`, write a lemma applying to -- the more general `associatedHom` and place it in the previous section. + /-- `associated` is the linear map that sends a quadratic map over a commutative ring to its associated symmetric bilinear map. -/ abbrev associated : QuadraticMap R M N →ₗ[R] BilinMap R M N := @@ -935,17 +998,19 @@ theorem coe_associatedHom : open LinearMap in @[simp] -theorem associated_linMulLin (f g : M →ₗ[R] R) : - associated (R := R) (linMulLin f g) = +theorem associated_linMulLin [Invertible (2 : R)] (f g : M →ₗ[R] R) : + associated (R := R) (N := R) (linMulLin f g) = ⅟ (2 : R) • ((mul R R).compl₁₂ f g + (mul R R).compl₁₂ g f) := by ext simp only [associated_apply, linMulLin_apply, map_add, smul_add, LinearMap.add_apply, - LinearMap.smul_apply, compl₁₂_apply, mul_apply', smul_eq_mul] + LinearMap.smul_apply, compl₁₂_apply, mul_apply', smul_eq_mul, invOf_smul_eq_iff] + simp only [smul_add, LinearMap.smul_def, Module.End.ofNat_apply, nsmul_eq_mul, Nat.cast_ofNat, + mul_invOf_cancel_left'] ring_nf open LinearMap in @[simp] -lemma associated_sq : associated (R := R) sq = mul R R := +lemma associated_sq [Invertible (2 : R)] : associated (R := R) sq = mul R R := (associated_linMulLin (id) (id)).trans <| by simp only [smul_add, invOf_two_smul_add_invOf_two_smul]; rfl @@ -1296,7 +1361,7 @@ theorem weightedSumSquares_apply [Monoid S] [DistribMulAction S R] [SMulCommClas /-- On an orthogonal basis, the basis representation of `Q` is just a sum of squares. -/ theorem basisRepr_eq_of_iIsOrtho {R M} [CommRing R] [AddCommGroup M] [Module R M] - [Invertible (2 : R)] (Q : QuadraticMap R M R) (v : Basis ι R M) + [Invertible (2 : R)] (Q : QuadraticForm R M) (v : Basis ι R M) (hv₂ : (associated (R := R) Q).IsOrthoᵢ v) : Q.basisRepr v = weightedSumSquares _ fun i => Q (v i) := by ext w @@ -1304,7 +1369,7 @@ theorem basisRepr_eq_of_iIsOrtho {R M} [CommRing R] [AddCommGroup M] [Module R M refine sum_congr rfl fun j hj => ?_ rw [← @associated_eq_self_apply R, LinearMap.map_sum₂, sum_eq_single_of_mem j hj] · rw [LinearMap.map_smul, LinearMap.map_smul₂, smul_eq_mul, associated_apply, smul_eq_mul, - smul_eq_mul, smul_eq_mul] + smul_eq_mul, LinearMap.smul_def, half_moduleEnd_apply_eq_half_smul] ring_nf · intro i _ hij rw [LinearMap.map_smul, LinearMap.map_smul₂, hv₂ hij] diff --git a/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean b/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean index 48db5f20ea5c2..803a8340d71e3 100644 --- a/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean +++ b/Mathlib/LinearAlgebra/QuadraticForm/Dual.lean @@ -141,8 +141,11 @@ def toDualProd (Q : QuadraticForm R M) [Invertible (2 : R)] : LinearMap.coe_toAddHom, LinearMap.prod_apply, Pi.prod, LinearMap.add_apply, LinearMap.coe_comp, Function.comp_apply, LinearMap.fst_apply, LinearMap.snd_apply, LinearMap.sub_apply, dualProd_apply, polarBilin_apply_apply, prod_apply, neg_apply] - simp [polar_comm _ x.1 x.2, ← sub_add, mul_sub, sub_mul, smul_sub, Submonoid.smul_def, ← - sub_eq_add_neg (Q x.1) (Q x.2)] + simp only [polar_sub_right, polar_self, nsmul_eq_mul, Nat.cast_ofNat, polar_comm _ x.1 x.2, + smul_sub, LinearMap.smul_def, sub_add_sub_cancel, ← sub_eq_add_neg (Q x.1) (Q x.2)] + rw [← LinearMap.map_sub (⅟ 2 : Module.End R R), ← mul_sub, ← LinearMap.smul_def] + simp only [LinearMap.smul_def, half_moduleEnd_apply_eq_half_smul, smul_eq_mul, + invOf_mul_cancel_left'] /-! TODO: show that `QuadraticForm.toDualProd` is an `QuadraticForm.IsometryEquiv` From bcc3ed61725456f92fc9e6bd4fa23ef842176e77 Mon Sep 17 00:00:00 2001 From: yhtq <1414672068@qq.com> Date: Sat, 23 Nov 2024 08:16:22 +0000 Subject: [PATCH 062/250] feat: add Algebra.compHom and related lemma (#18404) Add some simple definition and lemmas to construct Algebra structure from Algebra on another equivalent ring. As preparation for the [following result](https://github.com/mbkybky/GaloisRamification/blob/1070c2f8c3359e891d29e30b3e7185d5abc78a86/GaloisRamification/ToMathlib/IsGalois.lean#L111): If `K` and `K'` are fraction rings of `A`, and `L` and `L'` are fraction rings of `B`, then `IsGalois K L` implies `IsGalois K' L'`. --- Mathlib/Algebra/Algebra/Defs.lean | 43 +++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/Mathlib/Algebra/Algebra/Defs.lean b/Mathlib/Algebra/Algebra/Defs.lean index f21cc3fe2a3ef..40ac1bbeaf52e 100644 --- a/Mathlib/Algebra/Algebra/Defs.lean +++ b/Mathlib/Algebra/Algebra/Defs.lean @@ -175,6 +175,19 @@ def RingHom.toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) smul_def' _ _ := rfl toRingHom := i +-- just simple lemmas for a declaration that is itself primed, no need for docstrings +set_option linter.docPrime false in +theorem RingHom.smul_toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) + (h : ∀ c x, i c * x = x * i c) (r : R) (s : S) : + let _ := RingHom.toAlgebra' i h + r • s = i r * s := rfl + +set_option linter.docPrime false in +theorem RingHom.algebraMap_toAlgebra' {R S} [CommSemiring R] [Semiring S] (i : R →+* S) + (h : ∀ c x, i c * x = x * i c) : + @algebraMap R S _ _ (i.toAlgebra' h) = i := + rfl + /-- Creating an algebra from a morphism to a commutative semiring. -/ def RingHom.toAlgebra {R S} [CommSemiring R] [CommSemiring S] (i : R →+* S) : Algebra R S := i.toAlgebra' fun _ => mul_comm _ @@ -311,6 +324,36 @@ section end +section compHom + +variable (A) (f : S →+* R) + +/-- +Compose an `Algebra` with a `RingHom`, with action `f s • m`. + +This is the algebra version of `Module.compHom`. +-/ +abbrev compHom : Algebra S A where + smul s a := f s • a + toRingHom := (algebraMap R A).comp f + commutes' _ _ := Algebra.commutes _ _ + smul_def' _ _ := Algebra.smul_def _ _ + +theorem compHom_smul_def (s : S) (x : A) : + letI := compHom A f + s • x = f s • x := rfl + +theorem compHom_algebraMap_eq : + letI := compHom A f + algebraMap S A = (algebraMap R A).comp f := rfl + +theorem compHom_algebraMap_apply (s : S) : + letI := compHom A f + algebraMap S A s = (algebraMap R A) (f s) := rfl + +end compHom + + variable (R A) /-- The canonical ring homomorphism `algebraMap R A : R →+* A` for any `R`-algebra `A`, From 8eedc61a030954e3dad5a06bd750b29f54709126 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Sat, 23 Nov 2024 08:16:23 +0000 Subject: [PATCH 063/250] chore(discover-lean-pr-testing): better tracing, better output formatting (#19392) --- .github/workflows/discover-lean-pr-testing.yml | 13 ++++++++----- 1 file changed, 8 insertions(+), 5 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 2bee99c8e421d..8210a3e66daab 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -74,9 +74,12 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" + echo "$PRS" + echo "====================" echo "$PRS" | tr ' ' '\n' > prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) echo "$MATCHING_BRANCHES" + echo "====================" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" @@ -89,25 +92,25 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES $BRANCH" + RELEVANT_BRANCHES="$RELEVANT_BRANCHES\n- $BRANCH" fi done # Output the relevant branches echo "'$RELEVANT_BRANCHES'" - echo "branches=$RELEVANT_BRANCHES" >> "$GITHUB_OUTPUT" + printf "branches<> "$GITHUB_OUTPUT" - name: Check if there are relevant branches id: check-branches run: | if [ -z "${{ steps.find-branches.outputs.branches }}" ]; then - echo "no_branches=true" >> "$GITHUB_ENV" + echo "branches_exist=false" >> "$GITHUB_ENV" else - echo "no_branches=false" >> "$GITHUB_ENV" + echo "branches_exist=true" >> "$GITHUB_ENV" fi - name: Send message on Zulip - if: env.no_branches == 'false' + if: env.branches_exist == 'true' uses: zulip/github-actions-zulip/send-message@v1 with: api-key: ${{ secrets.ZULIP_API_KEY }} From fe1540557f73ff6f7031474eb40642ce2a48a42d Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sat, 23 Nov 2024 08:34:52 +0000 Subject: [PATCH 064/250] chore: bump to nightly-2024-11-23 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index ae2a1d4650fc0..fdc71e836bafa 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-22 +leanprover/lean4:nightly-2024-11-23 From ba61bf3e0b903f7186a3861a6cf1d704a1bed7c0 Mon Sep 17 00:00:00 2001 From: Matthew Robert Ballard Date: Sat, 23 Nov 2024 08:46:42 +0000 Subject: [PATCH 065/250] chore(CategoryTheory.ConcreteCategory): inline `ConcreteCategory.forget` (#19217) We already make `forget` `reducible` and further declarations downstream of it `abbrev`'s. It makes sense to `inline` it also to avoid Lean noticing it. --- Mathlib/CategoryTheory/ConcreteCategory/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean index 38882d4e9f448..ebeea719fe2dc 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Basic.lean @@ -55,7 +55,7 @@ class ConcreteCategory (C : Type u) [Category.{v} C] where /-- That functor is faithful -/ [forget_faithful : forget.Faithful] -attribute [reducible] ConcreteCategory.forget +attribute [inline, reducible] ConcreteCategory.forget attribute [instance] ConcreteCategory.forget_faithful /-- The forgetful functor from a concrete category to `Type u`. -/ From 4fd77b8b24b94a4d2eb8e4409675e0542f8a415e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 09:01:50 +0000 Subject: [PATCH 066/250] chore: missing MvPolynomial.eval lemmas (#19356) Revealed by a new user question on [zulip](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Adding.20more.20examples.20to.20mathlib/near/483822621). Co-authored-by: Eric Wieser --- Mathlib/Algebra/MvPolynomial/Basic.lean | 29 +++++++++++++++++++++++++ 1 file changed, 29 insertions(+) diff --git a/Mathlib/Algebra/MvPolynomial/Basic.lean b/Mathlib/Algebra/MvPolynomial/Basic.lean index b73915560d966..b649bbdca1c39 100644 --- a/Mathlib/Algebra/MvPolynomial/Basic.lean +++ b/Mathlib/Algebra/MvPolynomial/Basic.lean @@ -919,6 +919,14 @@ theorem eval₂_C (a) : (C a).eval₂ f g = f a := by theorem eval₂_one : (1 : MvPolynomial σ R).eval₂ f g = 1 := (eval₂_C _ _ _).trans f.map_one +@[simp] theorem eval₂_natCast (n : Nat) : (n : MvPolynomial σ R).eval₂ f g = n := + (eval₂_C _ _ _).trans (map_natCast f n) + +-- See note [no_index around OfNat.ofNat] +@[simp] theorem eval₂_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval₂ f g = OfNat.ofNat n := + eval₂_natCast f g n + @[simp] theorem eval₂_X (n) : (X n).eval₂ f g = g n := by simp [eval₂_monomial, f.map_one, X, prod_single_index, pow_one] @@ -1069,6 +1077,11 @@ theorem eval_C : ∀ a, eval f (C a) = a := theorem eval_X : ∀ n, eval f (X n) = f n := eval₂_X _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] theorem eval_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).eval f = OfNat.ofNat n := + map_ofNat _ n + @[simp] theorem smul_eval (x) (p : MvPolynomial σ R) (s) : eval x (s • p) = s * eval x p := by rw [smul_eq_C_mul, (eval x).map_mul, eval_C] @@ -1129,6 +1142,11 @@ theorem map_monomial (s : σ →₀ ℕ) (a : R) : map f (monomial s a) = monomi theorem map_C : ∀ a : R, map f (C a : MvPolynomial σ R) = C (f a) := map_monomial _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] protected theorem map_ofNat (n : Nat) [n.AtLeastTwo] : + (no_index (OfNat.ofNat n) : MvPolynomial σ R).map f = OfNat.ofNat n := + _root_.map_ofNat _ _ + @[simp] theorem map_X : ∀ n : σ, map f (X n : MvPolynomial σ R) = X n := eval₂_X _ _ @@ -1344,6 +1362,11 @@ theorem aeval_X (s : σ) : aeval f (X s : MvPolynomial _ R) = f s := theorem aeval_C (r : R) : aeval f (C r) = algebraMap R S₁ r := eval₂_C _ _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] theorem aeval_ofNat (n : Nat) [n.AtLeastTwo] : + aeval f (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + map_ofNat _ _ + theorem aeval_unique (φ : MvPolynomial σ R →ₐ[R] S₁) : φ = aeval (φ ∘ X) := by ext i simp @@ -1475,6 +1498,12 @@ theorem aevalTower_X (i : σ) : aevalTower g y (X i) = y i := theorem aevalTower_C (x : R) : aevalTower g y (C x) = g x := eval₂_C _ _ _ +-- See note [no_index around OfNat.ofNat] +@[simp] +theorem aevalTower_ofNat (n : Nat) [n.AtLeastTwo] : + aevalTower g y (no_index (OfNat.ofNat n) : MvPolynomial σ R) = OfNat.ofNat n := + _root_.map_ofNat _ _ + @[simp] theorem aevalTower_comp_C : (aevalTower g y : MvPolynomial σ R →+* A).comp C = g := RingHom.ext <| aevalTower_C _ _ From 25a3b654353e3d52ffd0c255f313600fcea9e34a Mon Sep 17 00:00:00 2001 From: Christopher Hoskin Date: Sat, 23 Nov 2024 09:23:50 +0000 Subject: [PATCH 067/250] feature(Order/SupClosed): Add lemmas for inserting upper or lower bounds into a Sup/Inf closed set (#18740) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add lemmas for inserting upper or lower bounds into a Sup/Inf closed set. e.g. When `s` is sup closed, so is `s` with `⊤` or `⊥` appended. Similarly for `s` inf closed. Co-authored-by: Christopher Hoskin --- Mathlib/Order/SupClosed.lean | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/Mathlib/Order/SupClosed.lean b/Mathlib/Order/SupClosed.lean index 61ddd8850360c..84fdcccc76f9d 100644 --- a/Mathlib/Order/SupClosed.lean +++ b/Mathlib/Order/SupClosed.lean @@ -79,6 +79,17 @@ lemma supClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeSup (α i {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, SupClosed (t i)) : SupClosed (s.pi t) := fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) +lemma SupClosed.insert_upperBounds {s : Set α} {a : α} (hs : SupClosed s) (ha : a ∈ upperBounds s) : + SupClosed (insert a s) := by + rw [SupClosed] + aesop + +lemma SupClosed.insert_lowerBounds {s : Set α} {a : α} (h : SupClosed s) (ha : a ∈ lowerBounds s) : + SupClosed (insert a s) := by + rw [SupClosed] + have ha' : ∀ b ∈ s, a ≤ b := fun _ a ↦ ha a + aesop + end Set section Finset @@ -144,6 +155,17 @@ lemma infClosed_pi {ι : Type*} {α : ι → Type*} [∀ i, SemilatticeInf (α i {t : ∀ i, Set (α i)} (ht : ∀ i ∈ s, InfClosed (t i)) : InfClosed (s.pi t) := fun _a ha _b hb _i hi ↦ ht _ hi (ha _ hi) (hb _ hi) +lemma InfClosed.insert_upperBounds {s : Set α} {a : α} (hs : InfClosed s) (ha : a ∈ upperBounds s) : + InfClosed (insert a s) := by + rw [InfClosed] + have ha' : ∀ b ∈ s, b ≤ a := fun _ a ↦ ha a + aesop + +lemma InfClosed.insert_lowerBounds {s : Set α} {a : α} (h : InfClosed s) (ha : a ∈ lowerBounds s) : + InfClosed (insert a s) := by + rw [InfClosed] + aesop + end Set section Finset From 9df30fc4ef0f92d820515f323987c4f368f80c3d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 20:36:21 +1100 Subject: [PATCH 068/250] ugh --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Group/AddChar.lean | 2 +- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 21 ++++- Mathlib/Algebra/Order/Field/Power.lean | 4 +- Mathlib/Analysis/Complex/Circle.lean | 6 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 6 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/Limits/Creates.lean | 12 +-- .../Limits/FunctorCategory/Basic.lean | 8 +- Mathlib/Data/Complex/Abs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/LazyList/Basic.lean | 10 +- Mathlib/Data/List/Indexes.lean | 12 +-- Mathlib/Data/List/Permutation.lean | 2 +- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Deprecated/AlgebraClasses.lean | 46 ++++----- Mathlib/Deprecated/ByteArray.lean | 10 +- Mathlib/Deprecated/Combinator.lean | 9 +- Mathlib/Deprecated/Equiv.lean | 4 +- Mathlib/Deprecated/Logic.lean | 74 +++++++-------- Mathlib/Deprecated/NatLemmas.lean | 8 +- Mathlib/Deprecated/RelClasses.lean | 10 +- Mathlib/GroupTheory/Coset/Defs.lean | 9 +- Mathlib/GroupTheory/Exponent.lean | 3 +- Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 10 +- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 20 ++-- Mathlib/LinearAlgebra/Prod.lean | 28 +++--- Mathlib/MeasureTheory/Function/L1Space.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Measure/WithDensityFinite.lean | 2 +- Mathlib/NumberTheory/ArithmeticFunction.lean | 3 +- Mathlib/NumberTheory/MulChar/Basic.lean | 6 +- Mathlib/Order/OmegaCompletePartialOrder.lean | 2 +- Mathlib/SetTheory/Cardinal/Aleph.lean | 18 ++-- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 94 +++++++++---------- Mathlib/SetTheory/Ordinal/Basic.lean | 8 +- Mathlib/SetTheory/Ordinal/Enum.lean | 2 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 66 ++++++------- Mathlib/SetTheory/Ordinal/Principal.lean | 4 +- Mathlib/SetTheory/ZFC/Basic.lean | 43 ++++----- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 3 +- Mathlib/Topology/Algebra/Group/Quotient.lean | 6 +- Mathlib/Topology/LocalAtTarget.lean | 4 +- Mathlib/Topology/MetricSpace/Polish.lean | 2 +- Mathlib/Topology/NoetherianSpace.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 2 +- 50 files changed, 315 insertions(+), 290 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 3dde47cda9794..7bc0dd9d0dc2e 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -634,7 +634,7 @@ end MonoidHom end MonoidHom set_option linter.deprecated false in -@[simp, deprecated (since := "2024-10-17")] +@[simp, deprecated "Deprecated without replacement." (since := "2024-10-17")] lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl namespace List diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 23505cff4f752..e66bd4f51f989 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -253,7 +253,7 @@ lemma ne_one_iff : ψ ≠ 1 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff lemma ne_zero_iff : ψ ≠ 0 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff /-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ -@[deprecated (since := "2024-06-06")] +@[deprecated "Deprecated without replacement." (since := "2024-06-06")] def IsNontrivial (ψ : AddChar A M) : Prop := ∃ a : A, ψ a ≠ 1 set_option linter.deprecated false in diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 80a073b792385..efbd286634e4a 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -32,7 +32,11 @@ lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by refine Cardinal.toNat_le_toNat Cardinal.mk_mul_le ?_ aesop (add simp [Cardinal.mul_lt_aleph0_iff, finite_mul]) -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_mul_le := natCard_mul_le +@[to_additive] alias card_mul_le := natCard_mul_le + +attribute [deprecated natCard_mul_le (since := "2024-09-30")] card_mul_le +attribute [deprecated natCard_add_le (since := "2024-09-30")] card_add_le + end Mul @@ -47,7 +51,10 @@ lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by rw [← image_inv, Nat.card_image_of_injective inv_injective] -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv +@[to_additive] alias card_inv := natCard_inv + +attribute [deprecated natCard_inv (since := "2024-09-30")] card_inv +attribute [deprecated natCard_neg (since := "2024-09-30")] card_neg end InvolutiveInv @@ -67,7 +74,10 @@ variable [Group G] {s t : Set G} lemma natCard_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by rw [div_eq_mul_inv, ← natCard_inv t]; exact natCard_mul_le -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_div_le := natCard_div_le +@[to_additive] alias card_div_le := natCard_div_le + +attribute [deprecated natCard_div_le (since := "2024-09-30")] card_div_le +attribute [deprecated natCard_sub_le (since := "2024-09-30")] card_sub_le variable [MulAction G α] @@ -79,8 +89,11 @@ lemma _root_.Cardinal.mk_smul_set (a : G) (s : Set α) : #↥(a • s) = #s := lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s := Nat.card_image_of_injective (MulAction.injective a) _ -@[to_additive (attr := deprecated (since := "2024-09-30"))] +@[to_additive] alias card_smul_set := Cardinal.mk_smul_set +attribute [deprecated Cardinal.mk_smul_set (since := "2024-09-30")] card_smul_set +attribute [deprecated Cardinal.mk_vadd_set (since := "2024-09-30")] card_vadd_set + end Group end Set diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index d363f2170a5c4..7df61ba43d34a 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -73,13 +73,13 @@ theorem zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : Injective (a ^ · : ℤ theorem zpow_inj (h₀ : 0 < a) (h₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := zpow_right_inj₀ h₀ h₁ -@[deprecated (since := "2024-10-08")] +@[deprecated "Deprecated without replacement." (since := "2024-10-08")] theorem zpow_le_max_of_min_le {x : α} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) := have : Antitone fun n : ℤ => x ^ (-n) := fun _ _ h => zpow_le_zpow_right₀ hx (neg_le_neg h) (this h).trans_eq this.map_min -@[deprecated (since := "2024-10-08")] +@[deprecated "Deprecated without replacement." (since := "2024-10-08")] theorem zpow_le_max_iff_min_le {x : α} (hx : 1 < x) {a b c : ℤ} : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) ↔ min a b ≤ c := by simp_rw [le_max_iff, min_le_iff, zpow_le_zpow_iff_right₀ hx, neg_le_neg_iff] diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index 72e66ae46c1b7..b56e79503ce1d 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -41,17 +41,17 @@ open ComplexConjugate /-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. Please use `Circle` when referring to the circle as a type. -/ -@[deprecated (since := "2024-07-24")] +@[deprecated "Deprecated without replacement." (since := "2024-07-24")] def circle : Submonoid ℂ := Submonoid.unitSphere ℂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "Deprecated without replacement." (since := "2024-07-24")] theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := mem_sphere_zero_iff_norm set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "Deprecated without replacement." (since := "2024-07-24")] theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by simp [Complex.abs, mem_circle_iff_abs] diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index dd86591976534..4f96b555b2c3c 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -92,7 +92,7 @@ lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := adj.leftAdjoint_preservesColimits @@ -216,7 +216,7 @@ lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := adj.rightAdjoint_preservesLimits @@ -240,7 +240,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "Deprecated without replacement." (since := "2024-11-18")] lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E := Functor.reflectsLimits_of_isEquivalence E diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 1c0f0936e00c6..a159d404af292 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -111,7 +111,7 @@ theorem comp_f {P Q R : Karoubi C} (f : P ⟶ Q) (g : Q ⟶ R) : (f ≫ g).f = f @[simp] theorem id_f {P : Karoubi C} : Hom.f (𝟙 P) = P.p := rfl -@[deprecated (since := "2024-07-15")] +@[deprecated "Deprecated without replacement." (since := "2024-07-15")] theorem id_eq {P : Karoubi C} : 𝟙 P = ⟨P.p, by repeat' rw [P.idem]⟩ := rfl /-- It is possible to coerce an object of `C` into an object of `Karoubi C`. diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 148eafbe8efa1..19ff0630c5b87 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -332,7 +332,7 @@ instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 _ _ @@ -342,7 +342,7 @@ lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) 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")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F := @@ -354,7 +354,7 @@ instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F := @@ -466,7 +466,7 @@ instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K (Cocones.functoriality K F).mapIso ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 _ _ @@ -477,7 +477,7 @@ instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F := @@ -489,7 +489,7 @@ instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (since := "2024-11-19")] lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F := diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 357916b1fe01e..a39f5182361aa 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -362,7 +362,7 @@ lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) exact isLimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 := @@ -374,7 +374,7 @@ lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Cate PreservesLimitsOfShape J F := ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 := @@ -387,7 +387,7 @@ lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) ⟨fun {L} _ => preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 := @@ -412,7 +412,7 @@ lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) exact isColimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "Deprecated without replacement." (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 diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 99319c9d298b3..acbd1c017c03a 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -199,7 +199,7 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated (since := "2024-02-14")] +@[deprecated "Deprecated without replacement." (since := "2024-02-14")] lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 21c9e27353985..de29683071348 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1115,7 +1115,7 @@ lemma succAbove_ne_last {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a ≠ last _) lemma succAbove_last_apply (i : Fin n) : succAbove (last n) i = castSucc i := by rw [succAbove_last] -@[deprecated (since := "2024-05-30")] +@[deprecated "Deprecated without replacement." (since := "2024-05-30")] lemma succAbove_lt_ge (p : Fin (n + 1)) (i : Fin n) : castSucc i < p ∨ p ≤ castSucc i := Nat.lt_or_ge (castSucc i) p diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Data/LazyList/Basic.lean index 7a9b820336505..00b68b137a00c 100644 --- a/Mathlib/Data/LazyList/Basic.lean +++ b/Mathlib/Data/LazyList/Basic.lean @@ -24,7 +24,7 @@ namespace LazyList open Function /-- Isomorphism between strict and lazy lists. -/ -@[deprecated (since := "2024-07-22")] +@[deprecated "Deprecated without replacement." (since := "2024-07-22")] def listEquivLazyList (α : Type*) : List α ≃ LazyList α where toFun := LazyList.ofList invFun := LazyList.toList @@ -39,12 +39,12 @@ def listEquivLazyList (α : Type*) : List α ≃ LazyList α where · simp [toList, ofList] · simpa [ofList, toList] -@[deprecated (since := "2024-07-22")] +@[deprecated "Deprecated without replacement." (since := "2024-07-22")] instance : Traversable LazyList where map := @LazyList.traverse Id _ traverse := @LazyList.traverse -@[deprecated (since := "2024-07-22")] +@[deprecated "Deprecated without replacement." (since := "2024-07-22")] instance : LawfulTraversable LazyList := by apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs · induction xs using LazyList.rec with @@ -71,7 +71,7 @@ instance : LawfulTraversable LazyList := by Function.comp_def, Thunk.pure, ofList] | mk _ ih => apply ih -@[deprecated (since := "2024-07-22"), simp] +@[deprecated "Deprecated without replacement." (since := "2024-07-22"), simp] theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with | nil => simp [LazyList.bind] @@ -81,7 +81,7 @@ theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by simp [ih] | mk f ih => simp_all -@[deprecated (since := "2024-07-22")] +@[deprecated "Deprecated without replacement." (since := "2024-07-22")] instance : LawfulMonad LazyList := LawfulMonad.mk' (id_map := by intro α xs diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 6c8ce5ff6c794..477ac3d84bfa8 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -84,7 +84,7 @@ theorem mapIdx_eq_ofFn (l : List α) (f : ℕ → α → β) : section deprecated /-- Lean3 `map_with_index` helper function -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β | _, [] => [] | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as @@ -92,12 +92,12 @@ protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List set_option linter.deprecated false in /-- Given a function `f : ℕ → α → β` and `as : List α`, `as = [a₀, a₁, ...]`, returns the list `[f 0 a₀, f 1 a₁, ...]`. -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := List.oldMapIdxCore f 0 as set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : ℕ) : l.oldMapIdxCore f n = l.oldMapIdx fun i a ↦ f (i + n) a := by induction' l with hd tl hl generalizing f n @@ -106,7 +106,7 @@ protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : simp only [List.oldMapIdxCore, hl, Nat.add_left_comm, Nat.add_comm, Nat.add_zero] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) (l₁ l₂ : List α), List.oldMapIdxCore f n (l₁ ++ l₂) = List.oldMapIdxCore f n l₁ ++ List.oldMapIdxCore f (n + l₁.length) l₂ := by @@ -134,7 +134,7 @@ protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) ( rw [Nat.add_assoc]; simp only [Nat.add_comm] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) (e : α), List.oldMapIdx f (l ++ [e]) = List.oldMapIdx f l ++ [f l.length e] := by intros f l e @@ -143,7 +143,7 @@ protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) ( simp only [Nat.zero_add]; rfl set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "Deprecated without replacement." (since := "2024-08-15")] protected theorem new_def_eq_old_def : ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by intro f diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 74a32e39c3263..849d50190c764 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -419,7 +419,7 @@ theorem length_permutations'Aux (s : List α) (x : α) : · simp · simpa using IH -@[deprecated (since := "2024-06-12")] +@[deprecated "Deprecated without replacement." (since := "2024-06-12")] theorem permutations'Aux_get_zero (s : List α) (x : α) (hn : 0 < length (permutations'Aux x s) := (by simp)) : (permutations'Aux x s).get ⟨0, hn⟩ = x :: s := diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 2e75e96025a44..d8394f09dc005 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -36,12 +36,12 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ -@[deprecated (since := "2024-08-29")] +@[deprecated "Deprecated without replacement." (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "Deprecated without replacement." (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index cc26f497422b2..d93789e6ec5dc 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -25,16 +25,16 @@ universe u v variable {α : Sort u} -@[deprecated (since := "2024-09-11")] +@[deprecated "Deprecated without replacement." (since := "2024-09-11")] class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where left_cancel : ∀ a b c, op a b = op a c → b = c -@[deprecated (since := "2024-09-11")] +@[deprecated "Deprecated without replacement." (since := "2024-09-11")] class IsRightCancel (α : Sort u) (op : α → α → α) : Prop where right_cancel : ∀ a b c, op a b = op c b → a = c /-- `IsTotalPreorder X r` means that the binary relation `r` on `X` is total and a preorder. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] class IsTotalPreorder (α : Sort u) (r : α → α → Prop) extends IsTrans α r, IsTotal α r : Prop /-- Every total pre-order is a pre-order. -/ @@ -45,11 +45,11 @@ instance (priority := 100) isTotalPreorder_isPreorder (α : Sort u) (r : α → /-- `IsIncompTrans X lt` means that for `lt` a binary relation on `X`, the incomparable relation `fun a b => ¬ lt a b ∧ ¬ lt b a` is transitive. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] class IsIncompTrans (α : Sort u) (lt : α → α → Prop) : Prop where incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeakOrder α lt] : IsIncompTrans α lt := { ‹IsStrictWeakOrder α lt› with } @@ -59,7 +59,7 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem incomp_trans [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := IsIncompTrans.incomp_trans _ _ _ @@ -68,7 +68,7 @@ section ExplicitRelationVariants variable (r) -@[elab_without_expected_type, deprecated (since := "2024-07-30")] +@[elab_without_expected_type, deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem incomp_trans_of [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := incomp_trans @@ -85,32 +85,32 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] def Equiv (a b : α) : Prop := ¬a ≺ b ∧ ¬b ≺ a local infixl:50 " ≈ " => @Equiv _ r -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem esymm {a b : α} : a ≈ b → b ≈ a := fun ⟨h₁, h₂⟩ => ⟨h₂, h₁⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem not_lt_of_equiv {a b : α} : a ≈ b → ¬a ≺ b := fun h => h.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem not_lt_of_equiv' {a b : α} : a ≈ b → ¬b ≺ a := fun h => h.2 variable [IsStrictWeakOrder α r] -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem erefl (a : α) : a ≈ a := ⟨irrefl a, irrefl a⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem etrans {a b c : α} : a ≈ b → b ≈ c → a ≈ c := incomp_trans -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] instance isEquiv : IsEquiv α (@Equiv _ r) where refl := erefl trans _ _ _ := etrans @@ -123,7 +123,7 @@ notation:50 a " ≈[" lt "]" b:50 => @Equiv _ lt a b end StrictWeakOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem isStrictWeakOrder_of_isTotalPreorder {α : Sort u} {le : α → α → Prop} {lt : α → α → Prop} [DecidableRel le] [IsTotalPreorder α le] (h : ∀ a b, lt a b ↔ ¬le b a) : IsStrictWeakOrder α lt := @@ -149,20 +149,20 @@ section LinearOrder variable {α : Type*} [LinearOrder α] set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] instance : IsTotalPreorder α (· ≤ ·) where trans := @le_trans _ _ total := le_total set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] instance isStrictWeakOrder_of_linearOrder : IsStrictWeakOrder α (· < ·) := have : IsTotalPreorder α (· ≤ ·) := by infer_instance -- Porting note: added isStrictWeakOrder_of_isTotalPreorder lt_iff_not_ge end LinearOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, lt a b → ¬lt b c ∧ ¬lt c b → lt a c := @fun a b c hab ⟨nbc, ncb⟩ => @@ -171,7 +171,7 @@ theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt a b ∧ ¬lt b a := incomp_trans_of lt ⟨nac, nca⟩ ⟨ncb, nbc⟩ absurd hab this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, ¬lt a b ∧ ¬lt b a → lt b c → lt a c := @fun a b c ⟨nab, nba⟩ hbc => @@ -180,7 +180,7 @@ theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt b c ∧ ¬lt c b := incomp_trans_of lt ⟨nba, nab⟩ ⟨nac, nca⟩ absurd hbc this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : ¬lt a b ∧ ¬lt b a → a = b := fun ⟨nab, nba⟩ => match trichotomous_of lt a b with @@ -188,17 +188,17 @@ theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α | Or.inr (Or.inl hab) => hab | Or.inr (Or.inr hba) => absurd hba nba -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem eq_of_eqv_lt {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : a ≈[lt]b → a = b := eq_of_incomp -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem incomp_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : ¬lt a b ∧ ¬lt b a ↔ a = b := Iff.intro eq_of_incomp fun hab => hab ▸ And.intro (irrefl_of lt a) (irrefl_of lt a) -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem eqv_lt_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : a ≈[lt]b ↔ a = b := incomp_iff_eq a b diff --git a/Mathlib/Deprecated/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean index c199be41b3f7c..35ac422df4fe2 100644 --- a/Mathlib/Deprecated/ByteArray.lean +++ b/Mathlib/Deprecated/ByteArray.lean @@ -19,7 +19,7 @@ set_option linter.deprecated false namespace Nat /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "Deprecated without replacement." (since := "2024-08-19")] def Up (ub a i : Nat) := i < a ∧ i < ub theorem Up.next {ub i} (h : i < ub) : Up ub (i+1) i := ⟨Nat.lt_succ_self _, h⟩ @@ -28,13 +28,13 @@ theorem Up.WF (ub) : WellFounded (Up ub) := Subrelation.wf (h₂ := (measure (ub - ·)).wf) fun ⟨ia, iu⟩ ↦ Nat.sub_lt_sub_left iu ia /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "Deprecated without replacement." (since := "2024-08-19")] def upRel (ub : Nat) : WellFoundedRelation Nat := ⟨Up ub, Up.WF ub⟩ end Nat /-- A terminal byte slice, a suffix of a byte array. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "Deprecated without replacement." (since := "2024-08-19")] structure ByteSliceT := (arr : ByteArray) (off : Nat) namespace ByteSliceT @@ -66,7 +66,7 @@ def toArray : ByteSlice → ByteArray universe u v /-- The inner loop of the `forIn` implementation for byte slices. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "Deprecated without replacement." (since := "2024-08-19")] def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β → m (ForInStep β)) (arr : ByteArray) (off _end : Nat) (i : Nat) (b : β) : m β := if h : i < _end then do @@ -75,7 +75,7 @@ def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β | ForInStep.yield b => have := Nat.Up.next h; loop f arr off _end (i+1) b else pure b -@[deprecated (since := "2024-08-19")] +@[deprecated "Deprecated without replacement." (since := "2024-08-19")] instance {m : Type u → Type v} : ForIn m ByteSlice UInt8 := ⟨fun ⟨arr, off, len⟩ b f ↦ forIn.loop f arr off (off + len) off b⟩ diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean index 5c79bf0837a59..3c16256244179 100644 --- a/Mathlib/Deprecated/Combinator.lean +++ b/Mathlib/Deprecated/Combinator.lean @@ -15,8 +15,11 @@ namespace Combinator universe u v w variable {α : Sort u} {β : Sort v} {γ : Sort w} -@[deprecated (since := "2024-07-27")] def I (a : α) := a -@[deprecated (since := "2024-07-27")] def K (a : α) (_b : β) := a -@[deprecated (since := "2024-07-27")] def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +def I (a : α) := a +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +def K (a : α) (_b : β) := a +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) end Combinator diff --git a/Mathlib/Deprecated/Equiv.lean b/Mathlib/Deprecated/Equiv.lean index 74bc9b68ecdd3..2491e31e194a1 100644 --- a/Mathlib/Deprecated/Equiv.lean +++ b/Mathlib/Deprecated/Equiv.lean @@ -21,10 +21,10 @@ variable {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α set_option linter.deprecated false -@[deprecated (since := "2024-09-11")] +@[deprecated "Deprecated without replacement." (since := "2024-09-11")] instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ -@[deprecated (since := "2024-09-11")] +@[deprecated "Deprecated without replacement." (since := "2024-09-11")] instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ diff --git a/Mathlib/Deprecated/Logic.lean b/Mathlib/Deprecated/Logic.lean index f88642ec745a7..ce91b89b2cdf4 100644 --- a/Mathlib/Deprecated/Logic.lean +++ b/Mathlib/Deprecated/Logic.lean @@ -37,43 +37,43 @@ local infix:65 (priority := high) " + " => g def Commutative := ∀ a b, a * b = b * a @[deprecated Std.Associative (since := "2024-09-13")] def Associative := ∀ a b c, (a * b) * c = a * (b * c) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def LeftIdentity := ∀ a, one * a = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def RightIdentity := ∀ a, a * one = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def RightInverse := ∀ a, a * a⁻¹ = one -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def LeftCancelative := ∀ a b c, a * b = a * c → b = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def RightCancelative := ∀ a b c, a * b = c * b → a = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def LeftDistributive := ∀ a b c, a * (b + c) = a * b + a * c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def RightDistributive := ∀ a b c, (a + b) * c = a * c + b * c end Binary @[deprecated (since := "2024-09-03")] alias not_of_eq_false := of_eq_false -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem cast_proof_irrel {β : Sort u} (h₁ h₂ : α = β) (a : α) : cast h₁ a = cast h₂ a := rfl @[deprecated (since := "2024-09-03")] alias eq_rec_heq := eqRec_heq @[deprecated (since := "2024-09-03")] alias heq_prop := proof_irrel_heq -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_left {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a = a') → (h₂ : Eq.rec (motive := fun a _ ↦ φ a) p₁ e = p₂) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_right {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a' = a) → (h₂ : p₁ = Eq.rec (motive := fun a _ ↦ φ a) p₂ e) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem of_heq_true {a : Prop} (h : HEq a True) : a := of_eq_true (eq_of_heq h) -@[deprecated (since := "2024-09-03")] +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] theorem eq_rec_compose {α β φ : Sort u} : ∀ (p₁ : β = φ) (p₂ : α = β) (a : α), (Eq.recOn p₁ (Eq.recOn p₂ a : β) : φ) = Eq.recOn (Eq.trans p₂ p₁) a @@ -114,20 +114,20 @@ theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) /- decidable -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem decide_True' (h : Decidable True) : decide True = true := by simp -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem decide_False' (h : Decidable False) : decide False = false := by simp namespace Decidable -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def recOn_true [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : p) (h₄ : h₁ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isTrue _ => rfl) h₄ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : ¬p) (h₄ : h₂ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isFalse _ => rfl) h₄ @@ -145,25 +145,25 @@ end Decidable @[deprecated (since := "2024-09-03")] alias decidableTrue := instDecidableTrue @[deprecated (since := "2024-09-03")] alias decidableFalse := instDecidableFalse -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def IsDecEq {α : Sort u} (p : α → α → Bool) : Prop := ∀ ⦃x y : α⦄, p x y = true → x = y -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def IsDecRefl {α : Sort u} (p : α → α → Bool) : Prop := ∀ x, p x x = true -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def decidableEq_of_bool_pred {α : Sort u} {p : α → α → Bool} (h₁ : IsDecEq p) (h₂ : IsDecRefl p) : DecidableEq α | x, y => if hp : p x y = true then isTrue (h₁ hp) else isFalse (fun hxy : x = y ↦ absurd (h₂ y) (by rwa [hxy] at hp)) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inl_refl {α : Sort u} [h : DecidableEq α] (a : α) : h a a = isTrue (Eq.refl a) := match h a a with | isTrue _ => rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} (n : a ≠ b) : h a b = isFalse n := match h a b with @@ -171,7 +171,7 @@ theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} /- subsingleton -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} [h₃ : ∀ h : p, Subsingleton (h₁ h)] [h₄ : ∀ h : ¬p, Subsingleton (h₂ h)] : Subsingleton (Decidable.recOn h h₂ h₁) := @@ -179,15 +179,15 @@ theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h | isTrue h => h₃ h | isFalse h => h₄ h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_pos {c t e : Prop} [Decidable c] (h : ite c t e) (hc : c) : t := (if_pos hc ▸ h :) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_neg {c t e : Prop} [Decidable c] (h : ite c t e) (hnc : ¬c) : e := (if_neg hnc ▸ h :) -@[deprecated (since := "2024-09-11")] +@[deprecated "Deprecated without replacement." (since := "2024-09-11")] theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : Decidable c] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -199,7 +199,7 @@ theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : | isFalse h₁, isTrue h₂ => absurd h₂ (Iff.mp (not_congr h_c) h₁) | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : Decidable c] (h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) : ite b x y ↔ ite c u v := match dec_b, dec_c with @@ -209,12 +209,12 @@ theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : De | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -- @[congr] -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem if_congr_prop {b c x y u v : Prop} [Decidable b] [Decidable c] (h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) : ite b x y ↔ ite c u v := if_ctx_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : c → (x ↔ u)) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -222,7 +222,7 @@ theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c (h_e : ¬c → (y ↔ v)) : ite b x y ↔ @ite _ c (decidable_of_decidable_of_iff h_c) u v := if_ctx_congr_prop (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : x ↔ u) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -230,7 +230,7 @@ theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h (h_e : y ↔ v) : ite b x y ↔ (@ite _ c (decidable_of_decidable_of_iff h_c) u v) := if_ctx_simp_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -241,31 +241,31 @@ theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] dite b x y = @dite _ c (decidable_of_decidable_of_iff h_c) u v := dif_ctx_congr (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] def AsTrue (c : Prop) [Decidable c] : Prop := if c then True else False -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib def AsFalse (c : Prop) [Decidable c] : Prop := if c then False else True -@[deprecated (since := "2024-09-03")] +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] theorem AsTrue.get {c : Prop} [h₁ : Decidable c] (_ : AsTrue c) : c := match h₁ with | isTrue h_c => h_c /- Equalities for rewriting let-expressions -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_eq {α : Sort u} {β : Sort v} {a₁ a₂ : α} (b : α → β) (h : a₁ = a₂) : (let x : α := a₁; b x) = (let x : α := a₂; b x) := congrArg b h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_heq {α : Sort v} {β : α → Sort u} {a₁ a₂ : α} (b : ∀ x : α, β x) (h : a₁ = a₂) : HEq (let x : α := a₁; b x) (let x : α := a₂; b x) := by cases h; rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem let_body_eq {α : Sort v} {β : α → Sort u} (a : α) {b₁ b₂ : ∀ x : α, β x} (h : ∀ x, b₁ x = b₂ x) : (let x : α := a; b₁ x) = (let x : α := a; b₂ x) := by exact h _ ▸ rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib theorem let_eq {α : Sort v} {β : Sort u} {a₁ a₂ : α} {b₁ b₂ : α → β} (h₁ : a₁ = a₂) (h₂ : ∀ x, b₁ x = b₂ x) : (let x : α := a₁; b₁ x) = (let x : α := a₂; b₂ x) := by simp [h₁, h₂] diff --git a/Mathlib/Deprecated/NatLemmas.lean b/Mathlib/Deprecated/NatLemmas.lean index 37fd557b3fec3..fa59bd023418a 100644 --- a/Mathlib/Deprecated/NatLemmas.lean +++ b/Mathlib/Deprecated/NatLemmas.lean @@ -28,7 +28,7 @@ namespace Nat /-! successor and predecessor -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "Deprecated without replacement." (since := "2024-08-23")] def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ m → B) : B := by induction n with | zero => exact H1 rfl @@ -36,7 +36,7 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] theorem one_eq_succ_zero : 1 = succ 0 := rfl @@ -44,7 +44,7 @@ theorem one_eq_succ_zero : 1 = succ 0 := -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P (succ n) 0) (H3 : ∀ n m, P n m → P (succ n) (succ m)) : ∀ n m : ℕ, P n m | 0, _m => H1 _ @@ -55,7 +55,7 @@ def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] theorem cond_decide_mod_two (x : ℕ) [d : Decidable (x % 2 = 1)] : cond (@decide (x % 2 = 1) d) 1 0 = x % 2 := by simp only [cond_eq_if, decide_eq_true_eq] diff --git a/Mathlib/Deprecated/RelClasses.lean b/Mathlib/Deprecated/RelClasses.lean index 357b10283c2d5..b9f1af89148f5 100644 --- a/Mathlib/Deprecated/RelClasses.lean +++ b/Mathlib/Deprecated/RelClasses.lean @@ -27,12 +27,14 @@ variable {α : Type u} open Function -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where +@[deprecated "Deprecated without replacement." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where +@[deprecated "Deprecated without replacement." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where -@[deprecated (since := "2024-07-30")] +@[deprecated "Deprecated without replacement." (since := "2024-07-30")] instance [LinearOrder α] : IsIncompTrans α (· < ·) := by infer_instance diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index f43753df9ac06..50c360a3c7413 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -170,7 +170,8 @@ theorem induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (Quoti instance : Coe α (α ⧸ s) := ⟨mk⟩ -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias induction_on' := induction_on +@[to_additive (attr := deprecated (since := "2024-08-04"))] +alias induction_on' := induction_on @[to_additive (attr := simp)] theorem quotient_liftOn_mk {β} (f : α → β) (h) (x : α) : Quotient.liftOn' (x : α ⧸ s) f h = f x := @@ -194,7 +195,8 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := _ ↔ leftRel s a b := Quotient.eq'' _ ↔ _ := by rw [leftRel_apply] -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq +@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-04"))] +alias eq' := QuotientGroup.eq @[to_additive] theorem out_eq' (a : α ⧸ s) : mk a.out = a := @@ -209,7 +211,8 @@ variable (s) 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] +@[to_additive (attr := deprecated "Deprecated without replacement." (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/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 6ee49e4b8a6ac..1d8a38f4b7f43 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -644,7 +644,8 @@ lemma inv_eq_self_of_orderOf_eq_two {x : G} (hx : orderOf x = 2) : -- TODO: delete /-- Any group of exponent two is abelian. -/ -@[to_additive (attr := reducible, deprecated (since := "2024-02-17")) +@[to_additive (attr := reducible, + deprecated "Deprecated without replacement." (since := "2024-02-17")) "Any additive group of exponent two is abelian."] def instCommGroupOfExponentTwo (hG : Monoid.exponent G = 2) : CommGroup G where mul_comm := mul_comm_of_exponent_two hG diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index 0768933dedc46..006312d2cecef 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -58,7 +58,7 @@ namespace LinearMap namespace BilinForm -@[deprecated (since := "2024-04-14")] +@[deprecated "Deprecated without replacement." (since := "2024-04-14")] theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y' | _, _, _, _, rfl, rfl => rfl @@ -101,7 +101,7 @@ theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _ -@[deprecated (since := "2024-04-14")] +@[deprecated "Deprecated without replacement." (since := "2024-04-14")] theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := rfl @@ -111,7 +111,7 @@ theorem zero_apply (x y : M) : (0 : BilinForm R M) x y = 0 := variable (B D B₁ D₁) -@[deprecated (since := "2024-04-14")] +@[deprecated "Deprecated without replacement." (since := "2024-04-14")] theorem coe_add : ⇑(B + D) = B + D := rfl @@ -119,7 +119,7 @@ theorem coe_add : ⇑(B + D) = B + D := theorem add_apply (x y : M) : (B + D) x y = B x y + D x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "Deprecated without replacement." (since := "2024-04-14")] theorem coe_neg : ⇑(-B₁) = -B₁ := rfl @@ -127,7 +127,7 @@ theorem coe_neg : ⇑(-B₁) = -B₁ := theorem neg_apply (x y : M₁) : (-B₁) x y = -B₁ x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "Deprecated without replacement." (since := "2024-04-14")] theorem coe_sub : ⇑(B₁ - D₁) = B₁ - D₁ := rfl diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 9c011c36ed52b..8fe60684491f7 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -58,16 +58,16 @@ section ToLin' def toLinHomAux₁ (A : BilinForm R M) (x : M) : M →ₗ[R] R := A x /-- Auxiliary definition to define `toLinHom`; see below. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] def toLinHomAux₂ (A : BilinForm R M) : M →ₗ[R] M →ₗ[R] R := A /-- The linear map obtained from a `BilinForm` by fixing the left co-ordinate and evaluating in the right. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] def toLinHom : BilinForm R M →ₗ[R] M →ₗ[R] M →ₗ[R] R := LinearMap.id set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem toLin'_apply (A : BilinForm R M) (x : M) : toLinHom (M := M) A x = A x := rfl @@ -112,7 +112,7 @@ def LinearMap.toBilinAux (f : M →ₗ[R] M →ₗ[R] R) : BilinForm R M := f set_option linter.deprecated false in /-- Bilinear forms are linearly equivalent to maps with two arguments that are linear in both. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] R := { BilinForm.toLinHom with invFun := LinearMap.toBilinAux @@ -121,35 +121,35 @@ def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] set_option linter.deprecated false in /-- A map with two arguments that is linear in both is linearly equivalent to bilinear form. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] def LinearMap.toBilin : (M →ₗ[R] M →ₗ[R] R) ≃ₗ[R] BilinForm R M := BilinForm.toLin.symm -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem LinearMap.toBilinAux_eq (f : M →ₗ[R] M →ₗ[R] R) : LinearMap.toBilinAux f = f := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem LinearMap.toBilin_symm : (LinearMap.toBilin.symm : BilinForm R M ≃ₗ[R] _) = BilinForm.toLin := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem BilinForm.toLin_symm : (BilinForm.toLin.symm : _ ≃ₗ[R] BilinForm R M) = LinearMap.toBilin := LinearMap.toBilin.symm_symm set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem LinearMap.toBilin_apply (f : M →ₗ[R] M →ₗ[R] R) (x y : M) : toBilin f x y = f x y := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "Deprecated without replacement." (since := "2024-04-26")] theorem BilinForm.toLin_apply (x : M) : BilinForm.toLin B x = B x := rfl diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 4a19736a68ba4..3d4187a8405b4 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -834,17 +834,17 @@ set_option linter.deprecated false /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i /-- Auxiliary definition for `tunnel`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M | 0 => ⟨⊤, LinearEquiv.ofTop ⊤ rfl⟩ | n + 1 => @@ -855,7 +855,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M)ᵒᵈ := -- Note: the hint `(α := _)` had to be added in https://github.com/leanprover-community/mathlib4/pull/8386 ⟨fun n => OrderDual.toDual (α := Submodule R M) (tunnel' f i n).1, @@ -867,24 +867,24 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) /-- Each `tailing f i n` is a copy of `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by dsimp [tailing, tunnelAux] rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by rw [disjoint_iff] @@ -893,7 +893,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by @@ -902,19 +902,19 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv apply Submodule.map_subtype_le /-- The supremum of all the copies of `N` found inside the tunnel. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by induction' n with n ih @@ -926,7 +926,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel -@[deprecated (since := "2024-06-05")] +@[deprecated "Deprecated without replacement." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := Disjoint.mono_right (tailing_le_tunnel f i _) (tailings_disjoint_tunnel f i _) diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 25e17db9f409b..b003cc9a9db03 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -443,7 +443,7 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe /-- This lemma is a special case of `Integrable.of_finite`. -/ -- Eternal deprecation for discoverability, don't remove -@[deprecated Integrable.of_finite, nolint deprecatedNoSince] +@[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite @[deprecated (since := "2024-02-05")] alias integrable_of_fintype := Integrable.of_finite diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 7ab48f4222914..10d9db322ebca 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -567,7 +567,7 @@ instance isFiniteMeasure_sfiniteSeq [h : SFinite μ] (n : ℕ) : IsFiniteMeasure h.1.choose_spec.1 n set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "Deprecated without replacement." (since := "2024-10-11")] instance isFiniteMeasure_sFiniteSeq [SFinite μ] (n : ℕ) : IsFiniteMeasure (sFiniteSeq μ n) := isFiniteMeasure_sfiniteSeq n diff --git a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean index 21abd06cd6cee..2f78a1dd221a8 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean @@ -117,7 +117,7 @@ noncomputable def Measure.densityToFinite (μ : Measure α) [SFinite μ] (a : α μ.rnDeriv μ.toFinite a set_option linter.deprecated false in -@[deprecated (since := "2024-10-04")] +@[deprecated "Deprecated without replacement." (since := "2024-10-04")] lemma densityToFinite_def (μ : Measure α) [SFinite μ] : μ.densityToFinite = μ.rnDeriv μ.toFinite := rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 5c13f9d53ebe2..77a36164d0f5b 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -1291,7 +1291,8 @@ theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h -@[deprecated (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : +@[deprecated "Deprecated without replacement." (since := "2024-06-09")] +theorem card_divisors (n : ℕ) (hn : n ≠ 0) : #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 2df970973c59d..1c082292d1c8c 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -391,13 +391,13 @@ lemma ne_one_iff {χ : MulChar R R'} : χ ≠ 1 ↔ ∃ a : Rˣ, χ a ≠ 1 := b simp only [Ne, eq_one_iff, not_forall] /-- A multiplicative character is *nontrivial* if it takes a value `≠ 1` on a unit. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "Deprecated without replacement." (since := "2024-06-16")] def IsNontrivial (χ : MulChar R R') : Prop := ∃ a : Rˣ, χ a ≠ 1 set_option linter.deprecated false in /-- A multiplicative character is nontrivial iff it is not the trivial character. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "Deprecated without replacement." (since := "2024-06-16")] theorem isNontrivial_iff (χ : MulChar R R') : χ.IsNontrivial ↔ χ ≠ 1 := by simp only [IsNontrivial, Ne, MulChar.ext_iff, not_forall, one_apply_coe] @@ -573,7 +573,7 @@ theorem sum_eq_zero_of_ne_one [IsDomain R'] {χ : MulChar R R'} (hχ : χ ≠ 1) simpa only [Finset.mul_sum, ← map_mul] using b.mulLeft_bijective.sum_comp _ set_option linter.deprecated false in -@[deprecated (since := "2024-06-16")] +@[deprecated "Deprecated without replacement." (since := "2024-06-16")] lemma IsNontrivial.sum_eq_zero [IsDomain R'] {χ : MulChar R R'} (hχ : χ.IsNontrivial) : ∑ a, χ a = 0 := sum_eq_zero_of_ne_one ((isNontrivial_iff _).mp hχ) diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 309394381dfc2..2774cd126ed34 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -731,7 +731,7 @@ theorem apply_mono {f g : α →𝒄 β} {x y : α} (h₁ : f ≤ g) (h₂ : x OrderHom.apply_mono (show (f : α →o β) ≤ g from h₁) h₂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "Deprecated without replacement." (since := "2024-07-27")] theorem ite_continuous' {p : Prop} [hp : Decidable p] (f g : α → β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => if p then f x else g x := by split_ifs <;> simp [*] diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 3970c6fc6cb57..43f94a2400a2d 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -541,7 +541,7 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm -@[deprecated (since := "2024-08-28")] +@[deprecated "Deprecated without replacement." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl @@ -555,7 +555,7 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' -@[deprecated (since := "2024-08-28")] +@[deprecated "Deprecated without replacement." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl @@ -571,11 +571,11 @@ theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max -@[deprecated (since := "2024-08-28")] +@[deprecated "Deprecated without replacement." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c -@[deprecated (since := "2024-08-28")] +@[deprecated "Deprecated without replacement." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o @@ -608,7 +608,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias aleph'_omega := aleph'_omega0 /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @@ -637,7 +637,7 @@ theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := -- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. /-- Ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "Deprecated without replacement." (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -645,16 +645,16 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ -@[deprecated (since := "2024-09-24")] +@[deprecated "Deprecated without replacement." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ /-- Infinite ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "Deprecated without replacement." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded -@[deprecated (since := "2024-09-24")] +@[deprecated "Deprecated without replacement." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index b309bcdad858b..eab39bcb9515e 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -147,7 +147,7 @@ protected theorem eq : #α = #β ↔ Nonempty (α ≃ β) := Quotient.eq' /-- Avoid using `Quotient.mk` to construct a `Cardinal` directly -/ -@[deprecated (since := "2024-10-24")] +@[deprecated "Deprecated without replacement." (since := "2024-10-24")] theorem mk'_def (α : Type u) : @Eq Cardinal ⟦α⟧ #α := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 4024caa0c1f7a..b289d46418d33 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -114,7 +114,7 @@ def StrictOrder.cof (r : α → α → Prop) : Cardinal := Order.cof (swap rᶜ) /-- The set in the definition of `Order.StrictOrder.cof` is nonempty. -/ -@[deprecated (since := "2024-10-22")] +@[deprecated "Deprecated without replacement." (since := "2024-10-22")] theorem StrictOrder.cof_nonempty (r : α → α → Prop) [IsIrrefl α r] : { c | ∃ S : Set α, Unbounded r S ∧ #S = c }.Nonempty := @Order.cof_nonempty α _ (IsRefl.swap rᶜ) diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index d281b72b36a02..efff5957a1330 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -35,7 +35,7 @@ noncomputable def toPGame (o : Ordinal.{u}) : PGame.{u} := termination_by o decreasing_by exact ((enumIsoToType o).symm x).prop -@[deprecated (since := "2024-09-22")] +@[deprecated "Deprecated without replacement." (since := "2024-09-22")] theorem toPGame_def (o : Ordinal) : o.toPGame = ⟨o.toType, PEmpty, fun x => ((enumIsoToType o).symm x).val.toPGame, PEmpty.elim⟩ := by rw [toPGame] diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 6cc7dfe51df9e..ffc9423ec3d0f 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -605,14 +605,14 @@ theorem one_add_omega0 : 1 + ω = ω := by cases a <;> cases b <;> intro H <;> cases' H with _ _ H _ _ H <;> [exact H.elim; exact Nat.succ_pos _; exact Nat.succ_lt_succ H] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias one_add_omega := one_add_omega0 @[simp] theorem one_add_of_omega0_le {o} (h : ω ≤ o) : 1 + o = o := by rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias one_add_of_omega_le := one_add_of_omega0_le /-! ### Multiplication of ordinals -/ @@ -1185,7 +1185,7 @@ def sup {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal.{max u v} := iSup f set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "Deprecated without replacement." (since := "2024-08-27")] theorem sSup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sSup (Set.range f) = sup.{_, v} f := rfl @@ -1252,14 +1252,14 @@ protected theorem lt_iSup_iff {ι} {f : ι → Ordinal.{u}} {a : Ordinal.{u}} [S a < iSup f ↔ ∃ i, a < f i := lt_ciSup_iff' (bddAbove_of_small _) -@[deprecated (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff +@[deprecated "Deprecated without replacement." (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff set_option linter.deprecated false in @[deprecated Ordinal.lt_iSup (since := "2024-08-27")] theorem lt_sup {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < sup.{_, v} f ↔ ∃ i, a < f i := by simpa only [not_forall, not_le] using not_congr (@sup_le_iff.{_, v} _ f a) -@[deprecated (since := "2024-08-27")] +@[deprecated "Deprecated without replacement." (since := "2024-08-27")] theorem ne_iSup_iff_lt_iSup {ι : Type u} {f : ι → Ordinal.{max u v}} : (∀ i, f i ≠ iSup f) ↔ ∀ i, f i < iSup f := forall_congr' fun i => (Ordinal.le_iSup f i).lt_iff_ne.symm @@ -1377,7 +1377,7 @@ theorem unbounded_range_of_sup_ge {α β : Type u} (r : α → α → Prop) [IsW unbounded_range_of_le_iSup r f h set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "Deprecated without replacement." (since := "2024-08-27")] theorem le_sup_shrink_equiv {s : Set Ordinal.{u}} (hs : Small.{u} s) (a) (ha : a ∈ s) : a ≤ sup.{u, u} fun x => ((@equivShrink s hs).symm x).val := by convert le_sup.{u, u} (fun x => ((@equivShrink s hs).symm x).val) ((@equivShrink s hs) ⟨a, ha⟩) @@ -1422,7 +1422,7 @@ theorem IsNormal.apply_of_isLimit {f : Ordinal.{u} → Ordinal.{v}} (H : IsNorma rw [← H.map_iSup, ho.iSup_Iio] set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "Deprecated without replacement." (since := "2024-08-27")] theorem sup_eq_sSup {s : Set Ordinal.{u}} (hs : Small.{u} s) : (sup.{u, u} fun x => (@equivShrink s hs).symm x) = sSup s := let hs' := bddAbove_iff_small.2 hs @@ -1986,12 +1986,12 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : Deprecated. If you need this value explicitly, write it in terms of `iSup`. If you just want an upper bound for the image of `op`, use that `Iio a ×ˢ Iio b` is a small set. -/ -@[deprecated (since := "2024-10-11")] +@[deprecated "Deprecated without replacement." (since := "2024-10-11")] def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) : Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) -@[deprecated (since := "2024-10-11")] +@[deprecated "Deprecated without replacement." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} (ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by @@ -2012,34 +2012,34 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by refine mex_le_of_ne fun i hi => ?_ @@ -2072,18 +2072,18 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by convert (config := {transparency := .default}) @@ -2091,23 +2091,23 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := @@ -2165,7 +2165,7 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias one_add_nat_cast := one_add_natCast -- See note [no_index around OfNat.ofNat] @@ -2179,43 +2179,43 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n | 0 => by simp | n + 1 => by rw [Nat.mul_succ, Nat.cast_add, natCast_mul m n, Nat.cast_succ, mul_add_one] -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_mul := natCast_mul @[deprecated Nat.cast_le (since := "2024-10-17")] theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_le := natCast_le @[deprecated Nat.cast_inj (since := "2024-10-17")] theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_inj := natCast_inj @[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_lt := natCast_lt @[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero @[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero @[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_pos := natCast_pos @[simp, norm_cast] @@ -2226,7 +2226,7 @@ theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by · apply (add_left_cancel n).1 rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @[simp, norm_cast] @@ -2241,7 +2241,7 @@ theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_div := natCast_div @[simp, norm_cast] @@ -2249,7 +2249,7 @@ theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by rw [← add_left_cancel, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_mod := natCast_mod @[simp] @@ -2257,7 +2257,7 @@ theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n | 0 => by simp | n + 1 => by simp [lift_natCast n] -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias lift_nat_cast := lift_natCast -- See note [no_index around OfNat.ofNat] @@ -2292,13 +2292,13 @@ theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : theorem lt_omega0 {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by simp_rw [← Cardinal.ord_aleph0, Cardinal.lt_ord, lt_aleph0, card_eq_nat] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias lt_omega := lt_omega0 theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := lt_omega0.2 ⟨_, rfl⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias nat_lt_omega := nat_lt_omega0 theorem omega0_pos : 0 < ω := @@ -2307,12 +2307,12 @@ theorem omega0_pos : 0 < ω := theorem omega0_ne_zero : ω ≠ 0 := omega0_pos.ne' -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias omega_ne_zero := omega0_ne_zero theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega0 1 -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 theorem isLimit_omega0 : IsLimit ω := @@ -2320,10 +2320,10 @@ theorem isLimit_omega0 : IsLimit ω := let ⟨n, e⟩ := lt_omega0.1 h rw [e]; exact nat_lt_omega0 (n + 1)⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias omega_isLimit := isLimit_omega0 theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := @@ -2332,7 +2332,7 @@ theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := let ⟨n, e⟩ := lt_omega0.1 h rw [e, ← succ_le_iff]; exact H (n + 1)⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias omega_le := omega0_le @[simp] @@ -2344,7 +2344,7 @@ set_option linter.deprecated false in theorem sup_natCast : sup Nat.cast = ω := iSup_natCast -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias sup_nat_cast := sup_natCast theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o @@ -2354,7 +2354,7 @@ theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias omega_le_of_isLimit := omega0_le_of_isLimit theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by @@ -2372,7 +2372,7 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ intro e simp only [e, mul_zero] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias isLimit_iff_omega_dvd := isLimit_iff_omega0_dvd theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) @@ -2416,7 +2416,7 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias IsNormal.apply_omega := IsNormal.apply_omega0 @[simp] @@ -2460,7 +2460,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by · rw [ord_aleph0] exact Ordinal.isLimit_omega0 -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 5a31bfe919ee9..649cfe8761545 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -387,7 +387,7 @@ def typein (r : α → α → Prop) [IsWellOrder α r] : @PrincipalSeg α Ordina alias typein.principalSeg := typein set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "Deprecated without replacement." (since := "2024-10-09")] theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : (typein.principalSeg r : α → Ordinal) = typein r := rfl @@ -513,7 +513,7 @@ noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where right_inv _ := enum_typein _ _ map_rel_iff' := enum_le_enum' _ -@[deprecated (since := "2024-08-26")] +@[deprecated "Deprecated without replacement." (since := "2024-08-26")] alias enumIsoOut := enumIsoToType instance small_Iio (o : Ordinal.{u}) : Small.{u} (Iio o) := @@ -925,7 +925,7 @@ theorem card_succ (o : Ordinal) : card (succ o) = card o + 1 := by theorem natCast_succ (n : ℕ) : ↑n.succ = succ (n : Ordinal) := rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "Deprecated without replacement." (since := "2024-04-17")] alias nat_cast_succ := natCast_succ instance uniqueIioOne : Unique (Iio (1 : Ordinal)) where @@ -1213,7 +1213,7 @@ alias card_typein_out_lt := card_typein_toType_lt theorem mk_Iio_ord_toType {c : Cardinal} (i : c.ord.toType) : #(Iio i) < c := card_typein_toType_lt c i -@[deprecated (since := "2024-08-26")] +@[deprecated "Deprecated without replacement." (since := "2024-08-26")] alias mk_Iio_ord_out_α := mk_Iio_ord_toType theorem ord_injective : Injective ord := by diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean index 4e5d4d38a7cf6..04236bb473897 100644 --- a/Mathlib/SetTheory/Ordinal/Enum.lean +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -33,7 +33,7 @@ termination_by o variable {s : Set Ordinal.{u}} -@[deprecated (since := "2024-09-20")] +@[deprecated "Deprecated without replacement." (since := "2024-09-20")] theorem enumOrd_def (o : Ordinal.{u}) : enumOrd s o = sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) := by rw [enumOrd] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 36b50a052b83d..4152802449ec8 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -52,7 +52,7 @@ least `a`, and `Ordinal.nfpFamily_le_fp` shows this is the least ordinal with th def nfpFamily (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : Ordinal := ⨆ i, List.foldr f a i -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpFamily_eq_sup (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : nfpFamily f a = ⨆ i, List.foldr f a i := rfl @@ -238,48 +238,48 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := ⟨fun h => by @@ -287,19 +287,19 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by rw [← not_iff_not] push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -307,7 +307,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by refine ⟨fun h => ?_, fun h i hi => ?_⟩ @@ -316,13 +316,13 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : ¬ BddAbove (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [not_bddAbove_iff] @@ -347,12 +347,12 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := isNormal_derivFamily _ @@ -360,7 +360,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -368,7 +368,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by unfold derivBFamily @@ -378,7 +378,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by rw [← le_iff_derivBFamily H] @@ -387,7 +387,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : exact h i hi /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "Deprecated without replacement." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : derivBFamily.{u, v} o f = enumOrd (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_bfamily H)] @@ -428,7 +428,7 @@ theorem iSup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) exact Ordinal.le_iSup _ _ set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "Deprecated without replacement." (since := "2024-08-27")] theorem sup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : (sup fun n : ℕ => f^[n] a) = nfp f a := by refine le_antisymm ?_ (sup_le fun l => ?_) @@ -580,7 +580,7 @@ theorem nfp_add_eq_mul_omega0 {a b} (hba : b ≤ a * ω) : nfp (a + ·) b = a * exact nfp_monotone (isNormal_add_right a).monotone (Ordinal.zero_le b) · dsimp; rw [← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias nfp_add_eq_mul_omega := nfp_add_eq_mul_omega0 theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω ≤ b := by @@ -593,14 +593,14 @@ theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω nth_rw 1 [← this] rwa [← add_assoc, ← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias add_eq_right_iff_mul_omega_le := add_eq_right_iff_mul_omega0_le theorem add_le_right_iff_mul_omega0_le {a b : Ordinal} : a + b ≤ b ↔ a * ω ≤ b := by rw [← add_eq_right_iff_mul_omega0_le] exact (isNormal_add_right a).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias add_le_right_iff_mul_omega_le := add_le_right_iff_mul_omega0_le theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * ω + b := by @@ -612,7 +612,7 @@ theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * · rw [deriv_succ, h, add_succ] exact nfp_eq_self (add_eq_right_iff_mul_omega0_le.2 ((le_add_right _ _).trans (le_succ _))) -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias deriv_add_eq_mul_omega_add := deriv_add_eq_mul_omega0_add /-! ### Fixed points of multiplication -/ @@ -645,7 +645,7 @@ theorem nfp_mul_eq_opow_omega0 {a b : Ordinal} (hb : 0 < b) (hba : b ≤ a ^ ω) rw [← nfp_mul_one ha] exact nfp_monotone (isNormal_mul_right ha).monotone (one_le_iff_pos.2 hb) -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias nfp_mul_eq_opow_omega := nfp_mul_eq_opow_omega0 theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = b) : @@ -659,7 +659,7 @@ theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = rw [← Ne, ← one_le_iff_ne_zero] at hb exact nfp_le_fp (isNormal_mul_right ha).monotone hb (le_of_eq hab) -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias eq_zero_or_opow_omega_le_of_mul_eq_right := eq_zero_or_opow_omega0_le_of_mul_eq_right theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω ∣ b := by @@ -677,7 +677,7 @@ theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω cases' h with c hc rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias mul_eq_right_iff_opow_omega_dvd := mul_eq_right_iff_opow_omega0_dvd theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : @@ -685,7 +685,7 @@ theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : rw [← mul_eq_right_iff_opow_omega0_dvd] exact (isNormal_mul_right ha).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias mul_le_right_iff_opow_omega_dvd := mul_le_right_iff_opow_omega0_dvd theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) @@ -705,7 +705,7 @@ theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) rw [add_zero, mul_lt_mul_iff_left (opow_pos ω ha)] at this rwa [succ_le_iff] -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias nfp_mul_opow_omega_add := nfp_mul_opow_omega0_add theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : @@ -718,7 +718,7 @@ theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : · rw [deriv_succ, h] exact nfp_mul_opow_omega0_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) -@[deprecated (since := "2024-09-30")] +@[deprecated "Deprecated without replacement." (since := "2024-09-30")] alias deriv_mul_eq_opow_omega_mul := deriv_mul_eq_opow_omega0_mul end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5d6417f174042..5c0a4122cd004 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -129,7 +129,7 @@ theorem not_bddAbove_principal (op : Ordinal → Ordinal → Ordinal) : exact ((le_nfp _ _).trans (ha (principal_nfp_iSup op (succ a)))).not_lt (lt_succ a) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "Deprecated without replacement." (since := "2024-10-11")] theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : Principal op (nfp (fun o' => blsub₂.{u, u, u} o' o' (@fun a _ b _ => op a b)) o) := by intro a b ha hb @@ -147,7 +147,7 @@ theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordin exact lt_blsub₂ (@fun a _ b _ => op a b) hm (hn.trans_le h) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "Deprecated without replacement." (since := "2024-10-11")] theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : Set.Unbounded (· < ·) { o | Principal op o } := fun o => ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index c7a0c35cdfcef..724bc0096f422 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -447,12 +447,12 @@ set_option linter.deprecated false /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet.{u} n → Prop | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl @@ -460,46 +460,46 @@ theorem Arity.equiv_const {a : PSet.{u}} : /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 -@[deprecated (since := "2024-09-02"), refl] +@[deprecated "Deprecated without replacement." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c | 0, _, _, _, hab, hcb => PSet.Equiv.euc hab hcb | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) -@[deprecated (since := "2024-09-02"), symm] +@[deprecated "Deprecated without replacement." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc -@[deprecated (since := "2024-09-02"), trans] +@[deprecated "Deprecated without replacement." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -611,7 +611,7 @@ set_option linter.deprecated false namespace Resp /-- Helper function for `PSet.eval`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def evalAux : ∀ {n}, { f : Resp n → OfArity ZFSet.{u} ZFSet.{u} n // ∀ a b : Resp n, Resp.Equiv a b → f a = f b } @@ -626,11 +626,11 @@ def evalAux : evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := rfl @@ -640,24 +640,25 @@ end Resp /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) | mk (f) : Definable n (Resp.eval n f) -attribute [deprecated (since := "2024-09-02"), instance] Definable.mk +attribute [deprecated "Deprecated without replacement." (since := "2024-09-02"), instance] + Definable.mk /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ /-- Turns a definable function into a function that respects equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s | _, ⟨_⟩ => rfl @@ -670,7 +671,7 @@ open PSet ZFSet set_option linter.deprecated false in /-- All functions are classically definable. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "Deprecated without replacement." (since := "2024-09-02")] noncomputable def allDefinable : ∀ {n} (F : OfArity ZFSet ZFSet n), Definable n F | 0, F => let p := @Quotient.exists_rep PSet _ F @@ -706,7 +707,7 @@ theorem exact {x y : PSet} : mk x = mk y → PSet.Equiv x y := Quotient.exact set_option linter.deprecated false in -@[deprecated (since := "2024-09-02"), simp] +@[deprecated "Deprecated without replacement." (since := "2024-09-02"), simp] theorem eval_mk {n f x} : (@Resp.eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) (mk x) = Resp.eval n (Resp.f f x) := rfl diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 1da40dadc219d..dd9e33f4aa641 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -266,7 +266,7 @@ theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff -@[to_additive (attr := deprecated (since := "2024-08-06"))] +@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-06"))] alias smul_mem_nhds := smul_mem_nhds_smul @[to_additive (attr := simp)] diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 91b1a4ca3bcd4..f7373a1da6da0 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -65,7 +65,8 @@ homomorphisms. Deprecated and changed from a `class` to a `structure`. Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead. -/ -@[to_additive (attr := deprecated (since := "2024-10-08"))] +@[to_additive (attr := deprecated "Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead." + (since := "2024-10-08"))] structure ContinuousMonoidHomClass (A B : outParam Type*) [Monoid A] [Monoid B] [TopologicalSpace A] [TopologicalSpace B] [FunLike F A B] extends MonoidHomClass F A B, ContinuousMapClass F A B : Prop diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean index 4a2b490786409..ee71b9f91af12 100644 --- a/Mathlib/Topology/Algebra/Group/Quotient.lean +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -76,7 +76,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := QuotientGroup.isOpenQuotientMap_mk.locallyCompactSpace -@[to_additive (attr := deprecated (since := "2024-10-05"))] +@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-10-05"))] theorem continuous_smul₁ (x : G ⧸ N) : Continuous fun g : G => g • x := continuous_id.smul continuous_const @@ -101,7 +101,7 @@ instance instSecondCountableTopology [SecondCountableTopology G] : SecondCountableTopology (G ⧸ N) := ContinuousConstSMul.secondCountableTopology -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-05"))] theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := inferInstance @@ -117,7 +117,7 @@ instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where exact continuous_mk.comp continuous_mul continuous_inv := continuous_inv.quotient_map' _ -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-05"))] theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := instTopologicalGroup N diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 9d801889f8176..091dcd05966e4 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -82,7 +82,7 @@ theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : simpa [isClosed_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "Deprecated without replacement." (since := "2024-04-02")] theorem Set.restrictPreimage_isClosedMap (s : Set β) (H : IsClosedMap f) : IsClosedMap (s.restrictPreimage f) := H.restrictPreimage s @@ -94,7 +94,7 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : simpa [isOpen_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "Deprecated without replacement." (since := "2024-04-02")] theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index d1e851e3c7d44..373e5d313be0f 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,7 +102,7 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated (since := "2024-02-23")] +@[deprecated "Deprecated without replacement." (since := "2024-02-23")] theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance /-- A countable product of Polish spaces is Polish. -/ diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 64689fd04f3f0..be41137f7a7b4 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -98,7 +98,7 @@ theorem noetherianSpace_iff_isCompact : NoetherianSpace α ↔ ∀ s : Set α, I instance [NoetherianSpace α] : WellFoundedLT (Closeds α) := Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› -@[deprecated (since := "2024-10-07")] +@[deprecated "Deprecated without replacement." (since := "2024-10-07")] theorem NoetherianSpace.wellFounded_closeds [NoetherianSpace α] : WellFounded fun s t : Closeds α => s < t := wellFounded_lt diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 07a4ecd8f0010..6a0a5700cbe59 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -341,7 +341,7 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated (since := "2024-02-15")] +@[deprecated "Deprecated without replacement." (since := "2024-02-15")] lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a export ClosedIciTopology (isClosed_ge') From d1d52a625399fe9f91ef60a29818f4005053e70c Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Sat, 23 Nov 2024 09:53:38 +0000 Subject: [PATCH 069/250] chore: update Mathlib dependencies 2024-11-23 (#19396) 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 762bb1280e3ba..6e18dc8970e13 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "8b52587ff32e2e443cce6109b5305341289339e7", + "rev": "0dc51ac7947ff6aa2c16bcffb64c46c7149d1276", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 0b15a0dc917eb5bdc9450c801c99423944d6045e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 11:01:40 +0000 Subject: [PATCH 070/250] feat: independence of functions under conditioning (#17915) From PFR --- Mathlib/Analysis/SpecificLimits/Basic.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 10 ++++ Mathlib/MeasureTheory/Function/LpSpace.lean | 2 +- Mathlib/MeasureTheory/Integral/Bochner.lean | 2 +- .../Probability/ConditionalProbability.lean | 4 +- Mathlib/Probability/Independence/Basic.lean | 57 +++++++++++++++++++ Mathlib/Probability/Independence/Kernel.lean | 53 +++++++++++++++++ 7 files changed, 125 insertions(+), 5 deletions(-) diff --git a/Mathlib/Analysis/SpecificLimits/Basic.lean b/Mathlib/Analysis/SpecificLimits/Basic.lean index 3d759d9f3fb33..e9bbbad7a0c52 100644 --- a/Mathlib/Analysis/SpecificLimits/Basic.lean +++ b/Mathlib/Analysis/SpecificLimits/Basic.lean @@ -623,7 +623,7 @@ theorem tendsto_factorial_div_pow_self_atTop : refine (eventually_gt_atTop 0).mono fun n hn ↦ ?_ rcases Nat.exists_eq_succ_of_ne_zero hn.ne.symm with ⟨k, rfl⟩ rw [← prod_range_add_one_eq_factorial, pow_eq_prod_const, div_eq_mul_inv, ← inv_eq_one_div, - prod_natCast, Nat.cast_succ, ← prod_inv_distrib, ← prod_mul_distrib, + prod_natCast, Nat.cast_succ, ← Finset.prod_inv_distrib, ← prod_mul_distrib, Finset.prod_range_succ'] simp only [prod_range_succ', one_mul, Nat.cast_add, zero_add, Nat.cast_one] refine diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 36003f724d580..bccdde7f1f4ab 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -173,6 +173,16 @@ protected theorem inv_div {a b : ℝ≥0∞} (htop : b ≠ ∞ ∨ a ≠ ∞) (h rw [← ENNReal.inv_ne_top] at hzero rw [ENNReal.div_eq_inv_mul, ENNReal.div_eq_inv_mul, ENNReal.mul_inv htop hzero, mul_comm, inv_inv] +lemma prod_inv_distrib {ι : Type*} {f : ι → ℝ≥0∞} {s : Finset ι} + (hf : s.toSet.Pairwise fun i j ↦ f i ≠ 0 ∨ f j ≠ ∞) : (∏ i ∈ s, f i)⁻¹ = ∏ i ∈ s, (f i)⁻¹ := by + induction' s using Finset.cons_induction with i s hi ih + · simp + simp [← ih (hf.mono <| by simp)] + refine ENNReal.mul_inv (not_or_of_imp fun hi₀ ↦ prod_ne_top fun j hj ↦ ?_) + (not_or_of_imp fun hi₀ ↦ Finset.prod_ne_zero_iff.2 fun j hj ↦ ?_) + · exact imp_iff_not_or.2 (hf (by simp) (by simp [hj]) <| .symm <| ne_of_mem_of_not_mem hj hi) hi₀ + · exact imp_iff_not_or.2 (hf (by simp [hj]) (by simp) <| ne_of_mem_of_not_mem hj hi).symm hi₀ + protected theorem mul_div_mul_left (a b : ℝ≥0∞) (hc : c ≠ 0) (hc' : c ≠ ⊤) : c * a / (c * b) = a / b := by rw [div_eq_mul_inv, div_eq_mul_inv, ENNReal.mul_inv (Or.inl hc) (Or.inl hc'), mul_mul_mul_comm, diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 5b9c9290dac61..582ddfc879f14 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -1192,7 +1192,7 @@ def compLpₗ (L : E →L[𝕜] F) : Lp E p μ →ₗ[𝕜] Lp F p μ where ext1 filter_upwards [Lp.coeFn_smul c f, coeFn_compLp L (c • f), Lp.coeFn_smul c (L.compLp f), coeFn_compLp L f] with _ ha1 ha2 ha3 ha4 - simp only [ha1, ha2, ha3, ha4, map_smul, Pi.smul_apply] + simp only [ha1, ha2, ha3, ha4, _root_.map_smul, Pi.smul_apply] /-- Composing `f : Lp E p μ` with `L : E →L[𝕜] F`, seen as a continuous `𝕜`-linear map on `Lp E p μ`. See also the similar diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index c85ae9df03c29..2cf72496737ad 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -638,7 +638,7 @@ theorem integral_sub (f g : α →₁[μ] E) : integral (f - g) = integral f - i theorem integral_smul (c : 𝕜) (f : α →₁[μ] E) : integral (c • f) = c • integral f := by simp only [integral] show (integralCLM' (E := E) 𝕜) (c • f) = c • (integralCLM' (E := E) 𝕜) f - exact map_smul (integralCLM' (E := E) 𝕜) c f + exact _root_.map_smul (integralCLM' (E := E) 𝕜) c f local notation "Integral" => @integralCLM α E _ _ μ _ _ diff --git a/Mathlib/Probability/ConditionalProbability.lean b/Mathlib/Probability/ConditionalProbability.lean index 6dcf1ba7c951b..4e4d89fc5ac91 100644 --- a/Mathlib/Probability/ConditionalProbability.lean +++ b/Mathlib/Probability/ConditionalProbability.lean @@ -70,8 +70,8 @@ and scaled by the inverse of `μ s` (to make it a probability measure): def cond (s : Set Ω) : Measure Ω := (μ s)⁻¹ • μ.restrict s -@[inherit_doc] scoped notation:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t @[inherit_doc] scoped notation:max μ "[|" s "]" => ProbabilityTheory.cond μ s +@[inherit_doc cond] scoped notation3:max μ "[" t " | " s "]" => ProbabilityTheory.cond μ s t /-- The conditional probability measure of measure `μ` on `{ω | X ω ∈ s}`. @@ -172,7 +172,7 @@ theorem inter_pos_of_cond_ne_zero (hms : MeasurableSet s) (hcst : μ[t|s] ≠ 0) simp [hms, Set.inter_comm, cond] lemma cond_pos_of_inter_ne_zero [IsFiniteMeasure μ] (hms : MeasurableSet s) (hci : μ (s ∩ t) ≠ 0) : - 0 < μ[|s] t := by + 0 < μ[t | s] := by rw [cond_apply hms] refine ENNReal.mul_pos ?_ hci exact ENNReal.inv_ne_zero.mpr (measure_ne_top _ _) diff --git a/Mathlib/Probability/Independence/Basic.lean b/Mathlib/Probability/Independence/Basic.lean index cfa5d0c19e7b7..a7ef0fcd188d8 100644 --- a/Mathlib/Probability/Independence/Basic.lean +++ b/Mathlib/Probability/Independence/Basic.lean @@ -721,4 +721,61 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β end IndepFun +variable {ι Ω α β : Type*} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α} + {mβ : MeasurableSpace β} {μ : Measure Ω} {X : ι → Ω → α} {Y : ι → Ω → β} {f : _ → Set Ω} + {t : ι → Set β} {s : Finset ι} + +/-- The probability of an intersection of preimages conditioning on another intersection factors +into a product. -/ +lemma cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ) + (hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i)) + (hy : ∀ i ∉ s, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + μ[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, μ[f i | Y i in t i] := by + have : IsProbabilityMeasure (μ : Measure Ω) := hindep.isProbabilityMeasure + classical + cases nonempty_fintype ι + let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i' + calc + _ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by + rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i) + _ = (μ (⋂ i, Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by + congr + calc + _ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by + congr + simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ] + _ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib] + _ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g] + _ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * μ (⋂ i, g i) := by + rw [hindep.meas_iInter] + exact fun i ↦ ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + _ = (∏ i, μ (Y i ⁻¹' t i))⁻¹ * ∏ i, μ (g i) := by + rw [hindep.meas_iInter] + intro i + by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g] + · obtain ⟨A, hA, hA'⟩ := hf i hi + exact .inter ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + ⟨A ×ˢ Set.univ, hA.prod .univ, by ext; simp [← hA']⟩ + · exact ⟨.univ ×ˢ t i, MeasurableSet.univ.prod (ht _), by ext; simp [eq_comm]⟩ + _ = ∏ i, (μ (Y i ⁻¹' t i))⁻¹ * μ (g i) := by + rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib] + exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _ + _ = ∏ i, if i ∈ s then μ[f i | Y i ⁻¹' t i] else 1 := by + refine Finset.prod_congr rfl fun i _ ↦ ?_ + by_cases hi : i ∈ s + · simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))] + · simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top μ _)] + _ = _ := by simp + +lemma iIndepFun.cond [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) μ) + (hy : ∀ i, μ (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + iIndepFun (fun _ ↦ mα) X μ[|⋂ i, Y i ⁻¹' t i] := by + rw [iIndepFun_iff] + intro s f hf + convert cond_iInter hY hindep hf (fun i _ ↦ hy _) ht using 2 with i hi + simpa using cond_iInter hY hindep (fun j hj ↦ hf _ <| Finset.mem_singleton.1 hj ▸ hi) + (fun i _ ↦ hy _) ht + end ProbabilityTheory diff --git a/Mathlib/Probability/Independence/Kernel.lean b/Mathlib/Probability/Independence/Kernel.lean index 85ed49143dcc1..3f4c1971fcafe 100644 --- a/Mathlib/Probability/Independence/Kernel.lean +++ b/Mathlib/Probability/Independence/Kernel.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne -/ import Mathlib.MeasureTheory.Constructions.Pi +import Mathlib.Probability.ConditionalProbability import Mathlib.Probability.Kernel.Basic import Mathlib.Tactic.Peel @@ -1212,4 +1213,56 @@ theorem iIndepSet.iIndepFun_indicator [Zero β] [One β] {m : MeasurableSpace β end IndepFun +variable {ι Ω α β : Type*} {mΩ : MeasurableSpace Ω} {mα : MeasurableSpace α} + {mβ : MeasurableSpace β} {κ : Kernel α Ω} {μ : Measure α} {X : ι → Ω → α} {Y : ι → Ω → β} + {f : _ → Set Ω} {t : ι → Set β} {s : Finset ι} + +/-- The probability of an intersection of preimages conditioning on another intersection factors +into a product. -/ +lemma iIndepFun.cond_iInter [Finite ι] (hY : ∀ i, Measurable (Y i)) + (hindep : iIndepFun (fun _ ↦ mα.prod mβ) (fun i ω ↦ (X i ω, Y i ω)) κ μ) + (hf : ∀ i ∈ s, MeasurableSet[mα.comap (X i)] (f i)) + (hy : ∀ᵐ a ∂μ, ∀ i ∉ s, κ a (Y i ⁻¹' t i) ≠ 0) (ht : ∀ i, MeasurableSet (t i)) : + ∀ᵐ a ∂μ, (κ a)[⋂ i ∈ s, f i | ⋂ i, Y i ⁻¹' t i] = ∏ i ∈ s, (κ a)[f i | Y i in t i] := by + classical + cases nonempty_fintype ι + let g (i' : ι) := if i' ∈ s then Y i' ⁻¹' t i' ∩ f i' else Y i' ⁻¹' t i' + have hYt i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (Y i ⁻¹' t i) := + ⟨.univ ×ˢ t i, .prod .univ (ht _), by ext; simp [eq_comm]⟩ + have hg i : MeasurableSet[(mα.prod mβ).comap fun ω ↦ (X i ω, Y i ω)] (g i) := by + by_cases hi : i ∈ s <;> simp only [hi, ↓reduceIte, g] + · obtain ⟨A, hA, hA'⟩ := hf i hi + exact (hYt _).inter ⟨A ×ˢ .univ, hA.prod .univ, by ext; simp [← hA']⟩ + · exact hYt _ + filter_upwards [hy, hindep.ae_isProbabilityMeasure, hindep.meas_iInter hYt, hindep.meas_iInter hg] + with a hy _ hYt hg + calc + _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a ((⋂ i, Y i ⁻¹' t i) ∩ ⋂ i ∈ s, f i) := by + rw [cond_apply]; exact .iInter fun i ↦ hY i (ht i) + _ = (κ a (⋂ i, Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by + congr + calc + _ = (⋂ i, Y i ⁻¹' t i) ∩ ⋂ i, if i ∈ s then f i else .univ := by + congr + simp only [Set.iInter_ite, Set.iInter_univ, Set.inter_univ] + _ = ⋂ i, Y i ⁻¹' t i ∩ (if i ∈ s then f i else .univ) := by rw [Set.iInter_inter_distrib] + _ = _ := Set.iInter_congr fun i ↦ by by_cases hi : i ∈ s <;> simp [hi, g] + _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * κ a (⋂ i, g i) := by + rw [hYt] + _ = (∏ i, κ a (Y i ⁻¹' t i))⁻¹ * ∏ i, κ a (g i) := by + rw [hg] + _ = ∏ i, (κ a (Y i ⁻¹' t i))⁻¹ * κ a (g i) := by + rw [Finset.prod_mul_distrib, ENNReal.prod_inv_distrib] + exact fun _ _ i _ _ ↦ .inr <| measure_ne_top _ _ + _ = ∏ i, if i ∈ s then (κ a)[f i | Y i ⁻¹' t i] else 1 := by + refine Finset.prod_congr rfl fun i _ ↦ ?_ + by_cases hi : i ∈ s + · simp only [hi, ↓reduceIte, g, cond_apply (hY i (ht i))] + · simp only [hi, ↓reduceIte, g, ENNReal.inv_mul_cancel (hy i hi) (measure_ne_top _ _)] + _ = _ := by simp + +-- TODO: We can't state `Kernel.iIndepFun.cond` (the `Kernel` analogue of +-- `ProbabilityTheory.iIndepFun.cond`) because we don't have a version of `ProbabilityTheory.cond` +-- for kernels + end ProbabilityTheory.Kernel From d1f4f68c3cd7162e8b86e5ae3a27eb0250fb8095 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 22:29:45 +1100 Subject: [PATCH 071/250] fix test --- MathlibTest/Simps.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/MathlibTest/Simps.lean b/MathlibTest/Simps.lean index 8def1aa96d6ef..2c3e0d2f69483 100644 --- a/MathlibTest/Simps.lean +++ b/MathlibTest/Simps.lean @@ -1203,7 +1203,7 @@ structure Foo where @[simps field field_2 field_9_fst] def myFoo : Foo := ⟨1, ⟨1, 1⟩, 1⟩ -structure Prod (X Y : Type _) extends Prod X Y +structure Prod (X Y : Type _) extends _root_.Prod X Y structure Prod2 (X Y : Type _) extends Prod X Y From 09a053e649888befa819426d00878d7f3aaca830 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 13:04:00 +0000 Subject: [PATCH 072/250] feat(NumberTheory/LSeries/PrimesInAP): von Mangoldt restricted to a residue class (#19368) This is the next step on the way to **Dirichlet's Theorem**. It defines `ArithmeticFunction.vonMangoldt.residueClass` as the function that is the von Mangoldt function on a given residue class and zero outside and proves some facts about it (in particular that it is a linear combination of twists of the von Mangoldt function by Dirichlet characters). See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483785853) on Zulip. The large import increase is caused by the need for more material to be able to state and prove the new things. --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 94 ++++++++++++++++++-- 1 file changed, 89 insertions(+), 5 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 0c1fc1a1ebf26..0740c4659750a 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -3,9 +3,10 @@ Copyright (c) 2024 Michael Stoll. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ -import Mathlib.Data.Nat.Factorization.PrimePow -import Mathlib.Topology.Algebra.InfiniteSum.Constructions -import Mathlib.Topology.Algebra.InfiniteSum.NatInt +import Mathlib.NumberTheory.DirichletCharacter.Orthogonality +import Mathlib.NumberTheory.LSeries.DirichletContinuation +import Mathlib.NumberTheory.LSeries.Linearity +import Mathlib.RingTheory.RootsOfUnity.AlgebraicallyClosed /-! # Dirichlet's Theorem on primes in arithmetic progression @@ -17,8 +18,8 @@ and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` This will be done in the following steps: 1. Some auxiliary lemmas on infinite products and sums -2. (TODO) Results on the von Mangoldt function restricted to a residue class -3. (TODO) Results on its L-series +2. Results on the von Mangoldt function restricted to a residue class +3. (TODO) Results on its L-series and an auxiliary function related to it 4. (TODO) Derivation of Dirichlet's Theorem -/ @@ -65,3 +66,86 @@ lemma tprod_eq_tprod_primes_mul_tprod_primes_of_mulSupport_subset_prime_powers { Function.Injective.prodMap (fun ⦃_ _⦄ a ↦ a) <| add_left_injective 1 end auxiliary + +/-! +### The L-series of the von Mangoldt function restricted to a residue class +-/ + +section arith_prog + +namespace ArithmeticFunction.vonMangoldt + +open Complex LSeries DirichletCharacter + +open scoped LSeries.notation + +variable {q : ℕ} (a : ZMod q) + +/-- The von Mangoldt function restricted to the residue class `a` mod `q`. -/ +noncomputable abbrev residueClass : ℕ → ℝ := + {n : ℕ | (n : ZMod q) = a}.indicator (vonMangoldt ·) + +lemma residueClass_nonneg (n : ℕ) : 0 ≤ residueClass a n := + Set.indicator_apply_nonneg fun _ ↦ vonMangoldt_nonneg + +lemma residueClass_le (n : ℕ) : residueClass a n ≤ vonMangoldt n := + Set.indicator_apply_le' (fun _ ↦ le_rfl) (fun _ ↦ vonMangoldt_nonneg) + +@[simp] +lemma residueClass_apply_zero : residueClass a 0 = 0 := by + simp only [Set.indicator_apply_eq_zero, Set.mem_setOf_eq, Nat.cast_zero, map_zero, ofReal_zero, + implies_true] + +lemma abscissaOfAbsConv_residueClass_le_one : + abscissaOfAbsConv ↗(residueClass a) ≤ 1 := by + refine abscissaOfAbsConv_le_of_forall_lt_LSeriesSummable fun y hy ↦ ?_ + unfold LSeriesSummable + have := LSeriesSummable_vonMangoldt <| show 1 < (y : ℂ).re by simp only [ofReal_re, hy] + convert this.indicator {n : ℕ | (n : ZMod q) = a} + ext1 n + by_cases hn : (n : ZMod q) = a + · simp +contextual only [term, Set.indicator, Set.mem_setOf_eq, hn, ↓reduceIte, apply_ite, + ite_self] + · simp +contextual only [term, Set.mem_setOf_eq, hn, not_false_eq_true, Set.indicator_of_not_mem, + ofReal_zero, zero_div, ite_self] + +variable [NeZero q] {a} + +/-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination +of twists of the von Mangoldt function with Dirichlet charaters. -/ +lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : + residueClass a n = + (q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * χ n * vonMangoldt n := by + rw [eq_inv_mul_iff_mul_eq₀ <| mod_cast (Nat.totient_pos.mpr q.pos_of_neZero).ne'] + simp +contextual only [residueClass, Set.indicator_apply, Set.mem_setOf_eq, apply_ite, + ofReal_zero, mul_zero, ← Finset.sum_mul, sum_char_inv_mul_char_eq ℂ ha n, eq_comm (a := a), + ite_mul, zero_mul, ↓reduceIte, ite_self] + +/-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination +of twists of the von Mangoldt function with Dirichlet charaters. -/ +lemma residueClass_eq (ha : IsUnit a) : + ↗(residueClass a) = (q.totient : ℂ)⁻¹ • + ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ • (fun n : ℕ ↦ χ n * vonMangoldt n) := by + ext1 n + simpa only [Pi.smul_apply, Finset.sum_apply, smul_eq_mul, ← mul_assoc] + using residueClass_apply ha n + +/-- The L-series of the von Mangoldt function restricted to the residue class `a` mod `q` +with `a` invertible in `ZMod q` is a linear combination of logarithmic derivatives of +L-functions of the Dirichlet characters mod `q` (on `re s > 1`). -/ +lemma LSeries_residueClass_eq (ha : IsUnit a) {s : ℂ} (hs : 1 < s.re) : + LSeries ↗(residueClass a) s = + -(q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * + (deriv (LFunction χ) s / LFunction χ s) := by + simp only [deriv_LFunction_eq_deriv_LSeries _ hs, LFunction_eq_LSeries _ hs, neg_mul, ← mul_neg, + ← Finset.sum_neg_distrib, ← neg_div, ← LSeries_twist_vonMangoldt_eq _ hs] + rw [eq_inv_mul_iff_mul_eq₀ <| mod_cast (Nat.totient_pos.mpr q.pos_of_neZero).ne'] + simp_rw [← LSeries_smul, + ← LSeries_sum <| fun χ _ ↦ (LSeriesSummable_twist_vonMangoldt χ hs).smul _] + refine LSeries_congr s fun {n} _ ↦ ?_ + simp only [Pi.smul_apply, residueClass_apply ha, smul_eq_mul, ← mul_assoc, + mul_inv_cancel_of_invertible, one_mul, Finset.sum_apply, Pi.mul_apply] + +end ArithmeticFunction.vonMangoldt + +end arith_prog From 64cf70aebf299bdf23a33ae80e1f1025193e0062 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 13:12:26 +0000 Subject: [PATCH 073/250] feat: `DecidableEq` instance for quotient groups (#19258) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Hard to believe this is missing, but you can check for yourself: ``` import Mathlib variable {α : Type*} [Group α] (s : Subgroup α) [DecidablePred (· ∈ s)] #synth DecidableEq (α ⧸ s) ``` From GrowthInGroups --- Mathlib/GroupTheory/Coset/Defs.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index f43753df9ac06..bd49ccf9a7568 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -89,6 +89,10 @@ instance leftRelDecidable [DecidablePred (· ∈ s)] : DecidableRel (leftRel s). instance instHasQuotientSubgroup : HasQuotient α (Subgroup α) := ⟨fun s => Quotient (leftRel s)⟩ +@[to_additive] +instance [DecidablePred (· ∈ s)] : DecidableEq (α ⧸ s) := + @Quotient.decidableEq _ _ (leftRelDecidable _) + /-- The equivalence relation corresponding to the partition of a group by right cosets of a subgroup. -/ @[to_additive "The equivalence relation corresponding to the partition of a group by right cosets From 142886c5f379dacb026bdc0fd88a4b6dbee48bec Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Sat, 23 Nov 2024 13:12:27 +0000 Subject: [PATCH 074/250] chore(RingTheory/PrincipalIdealDomain): update module docs (#19361) The module documentation for `RingTheory.PrincipalIdealDomain` was out of date, listing definitions that have been renamed and moved. --- Mathlib/RingTheory/PrincipalIdealDomain.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 7cde766527865..3a675457d9104 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -14,20 +14,21 @@ import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain A principal ideal ring (PIR) is a ring in which all left ideals are principal. A principal ideal domain (PID) is an integral domain which is a principal ideal ring. +The definition of `IsPrincipalIdealRing` can be found in `Mathlib.RingTheory.Ideal.Span`. + # Main definitions Note that for principal ideal domains, one should use `[IsDomain R] [IsPrincipalIdealRing R]`. There is no explicit definition of a PID. -Theorems about PID's are in the `principal_ideal_ring` namespace. +Theorems about PID's are in the `PrincipalIdealRing` namespace. -- `IsPrincipalIdealRing`: a predicate on rings, saying that every left ideal is principal. - `IsBezout`: the predicate saying that every finitely generated left ideal is principal. - `generator`: a generator of a principal ideal (or more generally submodule) - `to_uniqueFactorizationMonoid`: a PID is a unique factorization domain # Main results -- `to_maximal_ideal`: a non-zero prime ideal in a PID is maximal. +- `Ideal.IsPrime.to_maximal_ideal`: a non-zero prime ideal in a PID is maximal. - `EuclideanDomain.to_principal_ideal_domain` : a Euclidean domain is a PID. - `IsBezout.nonemptyGCDMonoid`: Every Bézout domain is a GCD domain. From 22a386f051ff3747f95dc507f0b80ed394ec9b2d Mon Sep 17 00:00:00 2001 From: Ahmad Alkhalawi Date: Sat, 23 Nov 2024 13:41:09 +0000 Subject: [PATCH 075/250] feat(RingTheory/DedekindDomain/IntegralClosure): Added `IsIntegralClosure.finite` (#18842) Let A be a Noetherian integrally closed domain with fraction field K. Let L/K be a finite separable field extension. Then the integral closure of A in L is finite over A. The proof immediately follows from `IsIntegralClosure.isNoetherian`. Co-authored-by: Ahmad Alkhalawi <46321632+4hma4d@users.noreply.github.com> --- Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean index 77679db0fe1c6..f38e19b607e41 100644 --- a/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean +++ b/Mathlib/RingTheory/DedekindDomain/IntegralClosure.lean @@ -185,6 +185,14 @@ theorem IsIntegralClosure.isNoetherianRing [IsIntegrallyClosed A] [IsNoetherianR IsNoetherianRing C := isNoetherianRing_iff.mpr <| isNoetherian_of_tower A (IsIntegralClosure.isNoetherian A K L C) +/-- If `L` is a finite separable extension of `K = Frac(A)`, where `A` is +integrally closed and Noetherian, the integral closure `C` of `A` in `L` is +finite over `A`. -/ +theorem IsIntegralClosure.finite [IsIntegrallyClosed A] [IsNoetherianRing A] : + Module.Finite A C := by + haveI := IsIntegralClosure.isNoetherian A K L C + exact Module.IsNoetherian.finite A C + /-- If `L` is a finite separable extension of `K = Frac(A)`, where `A` is a principal ring and `L` has no zero smul divisors by `A`, the integral closure `C` of `A` in `L` is a free `A`-module. -/ From 76b6113527563ce69dceb7d0d8f56bc2257843cf Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Sat, 23 Nov 2024 13:41:10 +0000 Subject: [PATCH 076/250] feat(Ergodic): add `zero_measure` constructors (#19045) --- Mathlib/Dynamics/Ergodic/Ergodic.lean | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/Mathlib/Dynamics/Ergodic/Ergodic.lean b/Mathlib/Dynamics/Ergodic/Ergodic.lean index b7bd2016eeee4..cc9abb51e7198 100644 --- a/Mathlib/Dynamics/Ergodic/Ergodic.lean +++ b/Mathlib/Dynamics/Ergodic/Ergodic.lean @@ -80,6 +80,9 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : PreErgodic f μ) (c : R) : PreErgodic f (c • μ) where aeconst_set _s hs hfs := (hf.aeconst_set hs hfs).anti <| ae_smul_measure_le _ +theorem zero_measure (f : α → α) : @PreErgodic α m f 0 where + aeconst_set _ _ _ := by simp + end PreErgodic namespace MeasureTheory.MeasurePreserving @@ -137,6 +140,11 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : QuasiErgodic f μ) (c : R) : QuasiErgodic f (c • μ) := ⟨hf.1.smul_measure _, hf.2.smul_measure _⟩ +theorem zero_measure {f : α → α} (hf : Measurable f) : @QuasiErgodic α m f 0 where + measurable := hf + absolutelyContinuous := by simp + toPreErgodic := .zero_measure f + end QuasiErgodic namespace Ergodic @@ -171,6 +179,11 @@ theorem smul_measure {R : Type*} [SMul R ℝ≥0∞] [IsScalarTower R ℝ≥0∞ (hf : Ergodic f μ) (c : R) : Ergodic f (c • μ) := ⟨hf.1.smul_measure _, hf.2.smul_measure _⟩ +theorem zero_measure {f : α → α} (hf : Measurable f) : @Ergodic α m f 0 where + measurable := hf + map_eq := by simp + toPreErgodic := .zero_measure f + section IsFiniteMeasure variable [IsFiniteMeasure μ] From b738c674ff23ff46cb9549010eb25b3a4c1fc79c Mon Sep 17 00:00:00 2001 From: Joseph Myers Date: Sat, 23 Nov 2024 13:41:11 +0000 Subject: [PATCH 077/250] =?UTF-8?q?feat(Data/Nat/Nth):=20`nth=20p=20n=20?= =?UTF-8?q?=E2=89=A0=200`=20lemmas=20(#19142)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Add a few small lemmas that are convenient when working with `Nat.nth`, when you have a hypothesis of the form `nth p n ≠ 0`. ```lean lemma lt_toFinset_card_of_nth_ne_zero {n : ℕ} (h : nth p n ≠ 0) (hf : (setOf p).Finite) : n < hf.toFinset.card := by ``` ```lean lemma nth_mem_of_ne_zero {n : ℕ} (h : nth p n ≠ 0) : p (Nat.nth p n) := ``` ```lean lemma nth_ne_zero_anti (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (hb : nth p b ≠ 0) : nth p a ≠ 0 := ``` --- Mathlib/Data/Nat/Nth.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/Mathlib/Data/Nat/Nth.lean b/Mathlib/Data/Nat/Nth.lean index 4cde210602f99..243d52d17d6f6 100644 --- a/Mathlib/Data/Nat/Nth.lean +++ b/Mathlib/Data/Nat/Nth.lean @@ -237,10 +237,21 @@ theorem nth_eq_zero {n} : · rintro (⟨h₀, rfl⟩ | ⟨hf, hle⟩) exacts [nth_zero_of_zero h₀, nth_of_card_le hf hle] +lemma lt_card_toFinset_of_nth_ne_zero {n : ℕ} (h : nth p n ≠ 0) (hf : (setOf p).Finite) : + n < #hf.toFinset := by + simp only [ne_eq, nth_eq_zero, not_or, not_exists, not_le] at h + exact h.2 hf + +lemma nth_mem_of_ne_zero {n : ℕ} (h : nth p n ≠ 0) : p (Nat.nth p n) := + nth_mem n (lt_card_toFinset_of_nth_ne_zero h) + theorem nth_eq_zero_mono (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (ha : nth p a = 0) : nth p b = 0 := by simp only [nth_eq_zero, h₀, false_and, false_or] at ha ⊢ exact ha.imp fun hf hle => hle.trans hab +lemma nth_ne_zero_anti (h₀ : ¬p 0) {a b : ℕ} (hab : a ≤ b) (hb : nth p b ≠ 0) : nth p a ≠ 0 := + mt (nth_eq_zero_mono h₀ hab) hb + theorem le_nth_of_lt_nth_succ {k a : ℕ} (h : a < nth p (k + 1)) (ha : p a) : a ≤ nth p k := by cases' (setOf p).finite_or_infinite with hf hf · rcases exists_lt_card_finite_nth_eq hf ha with ⟨n, hn, rfl⟩ From a4de385a76ac2e8eeb7aa73f47775fdadf38b287 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sat, 23 Nov 2024 13:41:12 +0000 Subject: [PATCH 078/250] feat(CategoryTheory): add some API for countable products and countable filtered colimits (#19192) --- .../Limits/Shapes/Countable.lean | 114 +++++++++++++++++- .../Category/LightProfinite/Basic.lean | 9 +- 2 files changed, 118 insertions(+), 5 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean index a50342b46d933..6f0db0a360ddb 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Countable.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.CategoryTheory.Limits.Final -import Mathlib.CategoryTheory.Limits.Shapes.FiniteLimits +import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts import Mathlib.CategoryTheory.Countable import Mathlib.Data.Countable.Defs /-! @@ -52,6 +52,29 @@ instance [Category.{v} J] [CountableCategory J] [HasCountableLimits C] : HasLimi have : HasLimitsOfShape (HomAsType J) C := HasCountableLimits.out (HomAsType J) hasLimitsOfShape_of_equivalence (homAsTypeEquiv J) +/-- A category has countable products if it has all products indexed by countable types. -/ +class HasCountableProducts where + out (J : Type) [Countable J] : HasProductsOfShape J C + +instance [HasCountableProducts C] : HasProductsOfShape J C := + have : Countable (Shrink.{0} J) := Countable.of_equiv _ (equivShrink.{0} J) + have : HasLimitsOfShape (Discrete (Shrink.{0} J)) C := HasCountableProducts.out _ + hasLimitsOfShape_of_equivalence (Discrete.equivalence (equivShrink.{0} J)).symm + +instance (priority := 100) hasCountableProducts_of_hasProducts [HasProducts C] : + HasCountableProducts C where + out _ := + have : HasProducts.{0} C := has_smallest_products_of_hasProducts + inferInstance + +instance (priority := 100) hasCountableProducts_of_hasCountableLimits [HasCountableLimits C] : + HasCountableProducts C where + out _ := inferInstance + +instance (priority := 100) hasFiniteProducts_of_hasCountableProducts [HasCountableProducts C] : + HasFiniteProducts C where + out _ := inferInstance + /-- A category has all countable colimits if every functor `J ⥤ C` with a `CountableCategory J` instance and `J : Type` has a colimit. @@ -74,8 +97,81 @@ instance [Category.{v} J] [CountableCategory J] [HasCountableColimits C] : HasCo have : HasColimitsOfShape (HomAsType J) C := HasCountableColimits.out (HomAsType J) hasColimitsOfShape_of_equivalence (homAsTypeEquiv J) +/-- A category has countable coproducts if it has all coproducts indexed by countable types. -/ +class HasCountableCoproducts where + out (J : Type) [Countable J] : HasCoproductsOfShape J C + +instance (priority := 100) hasCountableCoproducts_of_hasCoproducts [HasCoproducts C] : + HasCountableCoproducts C where + out _ := + have : HasCoproducts.{0} C := has_smallest_coproducts_of_hasCoproducts + inferInstance + +instance [HasCountableCoproducts C] : HasCoproductsOfShape J C := + have : Countable (Shrink.{0} J) := Countable.of_equiv _ (equivShrink.{0} J) + have : HasColimitsOfShape (Discrete (Shrink.{0} J)) C := HasCountableCoproducts.out _ + hasColimitsOfShape_of_equivalence (Discrete.equivalence (equivShrink.{0} J)).symm + +instance (priority := 100) hasCountableCoproducts_of_hasCountableColimits [HasCountableColimits C] : + HasCountableCoproducts C where + out _ := inferInstance + +instance (priority := 100) hasFiniteCoproducts_of_hasCountableCoproducts + [HasCountableCoproducts C] : HasFiniteCoproducts C where + out _ := inferInstance + section Preorder +namespace IsFiltered + +attribute [local instance] IsFiltered.nonempty + +variable {C} [Preorder J] [IsFiltered J] + +/-- The object part of the initial functor `ℕᵒᵖ ⥤ J` -/ +noncomputable def sequentialFunctor_obj : ℕ → J := fun + | .zero => (exists_surjective_nat _).choose 0 + | .succ n => (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose n) + (sequentialFunctor_obj n)).choose + +theorem sequentialFunctor_map : Monotone (sequentialFunctor_obj J) := + monotone_nat_of_le_succ fun n ↦ + leOfHom (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose n) + (sequentialFunctor_obj J n)).choose_spec.choose_spec.choose + +/-- +The initial functor `ℕᵒᵖ ⥤ J`, which allows us to turn cofiltered limits over countable preorders +into sequential limits. +-/ +noncomputable def sequentialFunctor : ℕ ⥤ J where + obj n := sequentialFunctor_obj J n + map h := homOfLE (sequentialFunctor_map J (leOfHom h)) + +theorem sequentialFunctor_final_aux (j : J) : ∃ (n : ℕ), j ≤ sequentialFunctor_obj J n := by + obtain ⟨m, h⟩ := (exists_surjective_nat _).choose_spec j + refine ⟨m + 1, ?_⟩ + simpa only [h] using leOfHom (IsFilteredOrEmpty.cocone_objs ((exists_surjective_nat _).choose m) + (sequentialFunctor_obj J m)).choose_spec.choose + +instance sequentialFunctor_final : (sequentialFunctor J).Final where + out d := by + obtain ⟨n, (g : d ≤ (sequentialFunctor J).obj n)⟩ := sequentialFunctor_final_aux J d + have : Nonempty (StructuredArrow d (sequentialFunctor J)) := + ⟨StructuredArrow.mk (homOfLE g)⟩ + apply isConnected_of_zigzag + refine fun i j ↦ ⟨[j], ?_⟩ + simp only [List.chain_cons, Zag, List.Chain.nil, and_true, ne_eq, not_false_eq_true, + List.getLast_cons, not_true_eq_false, List.getLast_singleton', reduceCtorEq] + clear! C + wlog h : j.right ≤ i.right + · exact or_comm.1 (this J d n g inferInstance j i (le_of_lt (not_le.mp h))) + · right + exact ⟨StructuredArrow.homMk (homOfLE h) rfl⟩ + +end IsFiltered + +namespace IsCofiltered + attribute [local instance] IsCofiltered.nonempty variable {C} [Preorder J] [IsCofiltered J] @@ -94,6 +190,9 @@ theorem sequentialFunctor_map : Antitone (sequentialFunctor_obj J) := /-- The initial functor `ℕᵒᵖ ⥤ J`, which allows us to turn cofiltered limits over countable preorders into sequential limits. + +TODO: redefine this as `(IsFiltered.sequentialFunctor Jᵒᵖ).leftOp`. This would need API for initial/ +final functors of the form `leftOp`/`rightOp`. -/ noncomputable def sequentialFunctor : ℕᵒᵖ ⥤ J where obj n := sequentialFunctor_obj J (unop n) @@ -153,6 +252,19 @@ For this we need to dualize this whole section. proof_wanted hasCountableColimits_of_hasFiniteColimits_and_hasSequentialColimits [HasFiniteColimits C] [HasLimitsOfShape ℕ C] : HasCountableColimits C +end IsCofiltered + end Preorder +@[deprecated (since := "2024-11-01")] alias sequentialFunctor := IsCofiltered.sequentialFunctor +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_obj := + IsCofiltered.sequentialFunctor_obj +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_map := + IsCofiltered.sequentialFunctor_map +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_initial_aux := + IsCofiltered.sequentialFunctor_initial_aux +@[deprecated (since := "2024-11-01")] alias sequentialFunctor_initial := + IsCofiltered.sequentialFunctor_initial +attribute [nolint defLemma] sequentialFunctor_initial + end CategoryTheory.Limits diff --git a/Mathlib/Topology/Category/LightProfinite/Basic.lean b/Mathlib/Topology/Category/LightProfinite/Basic.lean index 362a23b4df054..bd67bbffa19ea 100644 --- a/Mathlib/Topology/Category/LightProfinite/Basic.lean +++ b/Mathlib/Topology/Category/LightProfinite/Basic.lean @@ -297,10 +297,11 @@ instance instCountableDiscreteQuotient (S : LightProfinite) : /-- A profinite space which is light gives rise to a light profinite space. -/ noncomputable def toLightDiagram (S : LightProfinite.{u}) : LightDiagram.{u} where - diagram := sequentialFunctor _ ⋙ (lightToProfinite.obj S).fintypeDiagram - cone := (Functor.Initial.limitConeComp (sequentialFunctor _) (lightToProfinite.obj S).lim).cone - isLimit := - (Functor.Initial.limitConeComp (sequentialFunctor _) (lightToProfinite.obj S).lim).isLimit + diagram := IsCofiltered.sequentialFunctor _ ⋙ (lightToProfinite.obj S).fintypeDiagram + cone := (Functor.Initial.limitConeComp (IsCofiltered.sequentialFunctor _) + (lightToProfinite.obj S).lim).cone + isLimit := (Functor.Initial.limitConeComp (IsCofiltered.sequentialFunctor _) + (lightToProfinite.obj S).lim).isLimit end LightProfinite From 1b79c7d8d77de09fa646c73ea18316f7d4fe6575 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Sat, 23 Nov 2024 13:41:13 +0000 Subject: [PATCH 079/250] feat(CategoryTheory): add `Pi.map_isIso` and `hasProductsOfShape_of_hasProducts` + duals (#19195) --- Mathlib/CategoryTheory/Limits/Shapes/Products.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean index 0bf114008f136..ed9bd8a77cbda 100644 --- a/Mathlib/CategoryTheory/Limits/Shapes/Products.lean +++ b/Mathlib/CategoryTheory/Limits/Shapes/Products.lean @@ -373,6 +373,10 @@ from a family of isomorphisms between the factors. abbrev Pi.mapIso {f g : β → C} [HasProductsOfShape β C] (p : ∀ b, f b ≅ g b) : ∏ᶜ f ≅ ∏ᶜ g := lim.mapIso (Discrete.natIso fun X => p X.as) +instance Pi.map_isIso {f g : β → C} [HasProductsOfShape β C] (p : ∀ b, f b ⟶ g b) + [∀ b, IsIso <| p b] : IsIso <| Pi.map p := + inferInstanceAs (IsIso (Pi.mapIso (fun b ↦ asIso (p b))).hom) + section /- In this section, we provide some API for products when we are given a functor @@ -487,6 +491,10 @@ from a family of isomorphisms between the factors. abbrev Sigma.mapIso {f g : β → C} [HasCoproductsOfShape β C] (p : ∀ b, f b ≅ g b) : ∐ f ≅ ∐ g := colim.mapIso (Discrete.natIso fun X => p X.as) +instance Sigma.map_isIso {f g : β → C} [HasCoproductsOfShape β C] (p : ∀ b, f b ⟶ g b) + [∀ b, IsIso <| p b] : IsIso (Sigma.map p) := + inferInstanceAs (IsIso (Sigma.mapIso (fun b ↦ asIso (p b))).hom) + section /- In this section, we provide some API for coproducts when we are given a functor @@ -660,6 +668,12 @@ theorem hasProducts_of_limit_fans (lf : ∀ {J : Type w} (f : J → C), Fan f) ⟨(Cones.postcompose Discrete.natIsoFunctor.inv).obj (lf fun j => F.obj ⟨j⟩), (IsLimit.postcomposeInvEquiv _ _).symm (lf_isLimit _)⟩ } +instance (priority := 100) hasProductsOfShape_of_hasProducts [HasProducts.{w} C] (J : Type w) : + HasProductsOfShape J C := inferInstance + +instance (priority := 100) hasCoproductsOfShape_of_hasCoproducts [HasCoproducts.{w} C] + (J : Type w) : HasCoproductsOfShape J C := inferInstance + /-! (Co)products over a type with a unique term. -/ From 6f5a5765040cc6b1b7c36c3f80821448977a8899 Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 23 Nov 2024 13:41:14 +0000 Subject: [PATCH 080/250] feat: add "new file" markers to import reports (#19274) Flagging new files should help in identifying why some files appear to have increased the number of imports, when creating new files. This was suggested in #19194. --- scripts/import_trans_difference.sh | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/scripts/import_trans_difference.sh b/scripts/import_trans_difference.sh index 9c103e3eabdae..a13bebc4957cf 100755 --- a/scripts/import_trans_difference.sh +++ b/scripts/import_trans_difference.sh @@ -67,7 +67,16 @@ git checkout "${currCommit}" printf '\n\n
Import changes for all files\n\n%s\n\n
\n' "$( printf "|Files|Import difference|\n|-|-|\n" - (awk -F, -v all="${all}" -v ghLimit='261752' '{ diff[$1]+=$2 } END { + (awk -F, -v all="${all}" -v ghLimit='261752' -v newFiles="$( + # we pass the "A"dded files with respect to master, converting them to module names + git diff --name-only --diff-filter=A master | tr '\n' , | sed 's=\.lean,=,=g; s=/=.=g' + )" ' + BEGIN{ + # `arrayNewModules` maps integers to module names + split(newFiles, arrayNewModules, ",") + # `newModules` "just" stores the module names + for(v in arrayNewModules) { newModules[arrayNewModules[v]]=0 } + } { diff[$1]+=$2 } END { fileCount=0 outputLength=0 for(fil in diff) { @@ -75,7 +84,8 @@ printf '\n\n
Import changes for all files\n\n%s\n\n Date: Sat, 23 Nov 2024 13:41:16 +0000 Subject: [PATCH 081/250] chore: rename *of_smul_regular lemmas (#19326) --- Mathlib/LinearAlgebra/Dimension/Basic.lean | 5 ++++- Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean | 9 ++++++--- .../LinearAlgebra/RootSystem/Finite/Nondegenerate.lean | 2 +- 3 files changed, 11 insertions(+), 5 deletions(-) diff --git a/Mathlib/LinearAlgebra/Dimension/Basic.lean b/Mathlib/LinearAlgebra/Dimension/Basic.lean index 254ddee6fae30..16bbd326c3112 100644 --- a/Mathlib/LinearAlgebra/Dimension/Basic.lean +++ b/Mathlib/LinearAlgebra/Dimension/Basic.lean @@ -353,13 +353,16 @@ 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] +lemma rank_le_of_isSMulRegular {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) +@[deprecated (since := "2024-11-21")] +alias rank_le_of_smul_regular := rank_le_of_isSMulRegular + end end Module diff --git a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean index 0e7836ca78bb4..71be346b20eaa 100644 --- a/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean +++ b/Mathlib/LinearAlgebra/Dimension/StrongRankCondition.lean @@ -465,12 +465,15 @@ 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} +theorem LinearMap.finrank_le_of_isSMulRegular {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) ?_ + refine finrank_le_finrank_of_rank_le_rank (lift_le.mpr <| rank_le_of_isSMulRegular L L' hr h) ?_ rw [← Module.finrank_eq_rank R L'] exact nat_lt_aleph0 (finrank R ↥L') +@[deprecated (since := "2024-11-21")] +alias LinearMap.finrank_le_of_smul_regular := LinearMap.finrank_le_of_isSMulRegular + end StrongRankCondition diff --git a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean index 734860e88627e..1efe86c29f24f 100644 --- a/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean +++ b/Mathlib/LinearAlgebra/RootSystem/Finite/Nondegenerate.lean @@ -70,7 +70,7 @@ lemma finrank_rootSpan_map_polarization_eq_finrank_corootSpan : 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 + refine LinearMap.finrank_le_of_isSMulRegular 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 => ?_ From d6a85bb4f0e582e6f48ee7c49c1c6ce8601797f7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 14:14:52 +0000 Subject: [PATCH 082/250] feat(Polynomial): `p.natDegree = q.natDegree` if `p.degree = q.degree` (#19402) Also generalise `natDegree_lt_natDegree` to two rings From GrowthInGroups (LeanCamCombi) Co-authored-by: Andrew Yang --- Mathlib/Algebra/Polynomial/Degree/Operations.lean | 7 +++++-- 1 file changed, 5 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Polynomial/Degree/Operations.lean b/Mathlib/Algebra/Polynomial/Degree/Operations.lean index b64a85909e805..21d80bcb9d741 100644 --- a/Mathlib/Algebra/Polynomial/Degree/Operations.lean +++ b/Mathlib/Algebra/Polynomial/Degree/Operations.lean @@ -29,7 +29,7 @@ variable {R : Type u} {S : Type v} {a b c d : R} {n m : ℕ} section Semiring -variable [Semiring R] {p q r : R[X]} +variable [Semiring R] [Semiring S] {p q r : R[X]} theorem supDegree_eq_degree (p : R[X]) : p.toFinsupp.supDegree WithBot.some = p.degree := max_eq_sup_coe @@ -65,12 +65,15 @@ theorem natDegree_eq_of_le_of_coeff_ne_zero (pn : p.natDegree ≤ n) (p1 : p.coe p.natDegree = n := pn.antisymm (le_natDegree_of_ne_zero p1) -theorem natDegree_lt_natDegree {p q : R[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : +theorem natDegree_lt_natDegree {q : S[X]} (hp : p ≠ 0) (hpq : p.degree < q.degree) : p.natDegree < q.natDegree := by by_cases hq : q = 0 · exact (not_lt_bot <| hq ▸ hpq).elim rwa [degree_eq_natDegree hp, degree_eq_natDegree hq, Nat.cast_lt] at hpq +lemma natDegree_eq_natDegree {q : S[X]} (hpq : p.degree = q.degree) : + p.natDegree = q.natDegree := by simp [natDegree, hpq] + theorem coeff_eq_zero_of_degree_lt (h : degree p < n) : coeff p n = 0 := Classical.not_not.1 (mt le_degree_of_ne_zero (not_le_of_gt h)) From 2abe270b9a11bad57afac01e29a40b4e92e1fcc5 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 14:14:53 +0000 Subject: [PATCH 083/250] chore: rename `Set.nonempty_of_nonempty_subtype` to `Set.Nonempty.of_subtype` (#19408) This is shorter and allows anonymous dot notation. From GrowthInGroups (LeanCamCombi) --- Mathlib/Data/ENat/Lattice.lean | 2 +- Mathlib/Data/Set/Basic.lean | 7 ++++--- Mathlib/Data/Set/Image.lean | 2 +- Mathlib/GroupTheory/PGroup.lean | 2 +- 4 files changed, 7 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/ENat/Lattice.lean b/Mathlib/Data/ENat/Lattice.lean index d56040d7d9e21..fd0d60e48e317 100644 --- a/Mathlib/Data/ENat/Lattice.lean +++ b/Mathlib/Data/ENat/Lattice.lean @@ -106,7 +106,7 @@ lemma finite_of_sSup_lt_top (h : sSup s < ⊤) : s.Finite := by exact sSup_eq_top_of_infinite h lemma sSup_mem_of_nonempty_of_lt_top [Nonempty s] (hs' : sSup s < ⊤) : sSup s ∈ s := - Nonempty.csSup_mem nonempty_of_nonempty_subtype (finite_of_sSup_lt_top hs') + Nonempty.csSup_mem .of_subtype (finite_of_sSup_lt_top hs') lemma exists_eq_iSup_of_lt_top [Nonempty ι] (h : ⨆ i, f i < ⊤) : ∃ i, f i = ⨆ i, f i := diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index d5e0c9c761f33..8d28bac27f0b0 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -436,8 +436,9 @@ instance univ.nonempty [Nonempty α] : Nonempty (↥(Set.univ : Set α)) := instance instNonemptyTop [Nonempty α] : Nonempty (⊤ : Set α) := inferInstanceAs (Nonempty (univ : Set α)) -theorem nonempty_of_nonempty_subtype [Nonempty (↥s)] : s.Nonempty := - nonempty_subtype.mp ‹_› +theorem Nonempty.of_subtype [Nonempty (↥s)] : s.Nonempty := nonempty_subtype.mp ‹_› + +@[deprecated (since := "2024-11-23")] alias nonempty_of_nonempty_subtype := Nonempty.of_subtype /-! ### Lemmas about the empty set -/ @@ -1204,7 +1205,7 @@ theorem eq_empty_of_ssubset_singleton {s : Set α} {x : α} (hs : s ⊂ {x}) : s theorem eq_of_nonempty_of_subsingleton {α} [Subsingleton α] (s t : Set α) [Nonempty s] [Nonempty t] : s = t := - nonempty_of_nonempty_subtype.eq_univ.trans nonempty_of_nonempty_subtype.eq_univ.symm + Nonempty.of_subtype.eq_univ.trans Nonempty.of_subtype.eq_univ.symm theorem eq_of_nonempty_of_subsingleton' {α} [Subsingleton α] {s : Set α} (t : Set α) (hs : s.Nonempty) [Nonempty t] : s = t := diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 11aa60b7641a2..16c6dcc43e64f 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -408,7 +408,7 @@ theorem Nonempty.preimage {s : Set β} (hs : s.Nonempty) {f : α → β} (hf : S ⟨x, mem_preimage.2 <| hx.symm ▸ hy⟩ instance (f : α → β) (s : Set α) [Nonempty s] : Nonempty (f '' s) := - (Set.Nonempty.image f nonempty_of_nonempty_subtype).to_subtype + (Set.Nonempty.image f .of_subtype).to_subtype /-- image and preimage are a Galois connection -/ @[simp] diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 9b119d28f05df..071291b214ad5 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -189,7 +189,7 @@ theorem card_modEq_card_fixedPoints : Nat.card α ≡ Nat.card (fixedPoints G α theorem nonempty_fixed_point_of_prime_not_dvd_card (α) [MulAction G α] (hpα : ¬p ∣ Nat.card α) : (fixedPoints G α).Nonempty := have : Finite α := Nat.finite_of_card_ne_zero (fun h ↦ (h ▸ hpα) (dvd_zero p)) - @Set.nonempty_of_nonempty_subtype _ _ + @Set.Nonempty.of_subtype _ _ (by rw [← Finite.card_pos_iff, pos_iff_ne_zero] contrapose! hpα From 1620b76241d3341c82b959798ef4b6e44e1dc626 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 23 Nov 2024 14:58:15 +0000 Subject: [PATCH 084/250] feat(AlgebraicGeometry): preimmersions are stable under base change (#18915) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Cover/MorphismProperty.lean | 3 + .../Morphisms/ClosedImmersion.lean | 4 +- .../Morphisms/Preimmersion.lean | 40 ++-- .../Morphisms/SurjectiveOnStalks.lean | 186 ++++++++++++++++++ Mathlib/AlgebraicGeometry/Scheme.lean | 12 ++ Mathlib/Topology/LocalAtTarget.lean | 41 ++++ 7 files changed, 264 insertions(+), 23 deletions(-) create mode 100644 Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean diff --git a/Mathlib.lean b/Mathlib.lean index 48be73827fdd6..58d4cfdeb00c9 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -962,6 +962,7 @@ import Mathlib.AlgebraicGeometry.Morphisms.QuasiSeparated import Mathlib.AlgebraicGeometry.Morphisms.RingHomProperties import Mathlib.AlgebraicGeometry.Morphisms.Separated import Mathlib.AlgebraicGeometry.Morphisms.Smooth +import Mathlib.AlgebraicGeometry.Morphisms.SurjectiveOnStalks import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap import Mathlib.AlgebraicGeometry.Morphisms.UniversallyClosed import Mathlib.AlgebraicGeometry.Morphisms.UniversallyInjective diff --git a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean index e2226dd86717f..53a9c4dfd2092 100644 --- a/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean +++ b/Mathlib/AlgebraicGeometry/Cover/MorphismProperty.lean @@ -73,6 +73,9 @@ theorem Cover.iUnion_range {X : Scheme.{u}} (𝒰 : X.Cover P) : rw [Set.mem_iUnion] exact ⟨𝒰.f x, 𝒰.covers x⟩ +lemma Cover.exists_eq (𝒰 : X.Cover P) (x : X) : ∃ i y, (𝒰.map i).base y = x := + ⟨_, 𝒰.covers x⟩ + /-- Given a family of schemes with morphisms to `X` satisfying `P` that jointly cover `X`, this an associated `P`-cover of `X`. -/ @[simps] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index 1cb1635d9fcb0..d72195178231c 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -62,8 +62,8 @@ lemma eq_inf : @IsClosedImmersion = (topologically IsClosedEmbedding) ⊓ lemma iff_isPreimmersion {X Y : Scheme} {f : X ⟶ Y} : IsClosedImmersion f ↔ IsPreimmersion f ∧ IsClosed (Set.range f.base) := by - rw [and_comm, isClosedImmersion_iff, isPreimmersion_iff, ← and_assoc, isClosedEmbedding_iff, - @and_comm (IsEmbedding _)] + rw [isClosedImmersion_iff, isPreimmersion_iff, ← surjectiveOnStalks_iff, and_comm, and_assoc, + isClosedEmbedding_iff] lemma of_isPreimmersion {X Y : Scheme} (f : X ⟶ Y) [IsPreimmersion f] (hf : IsClosed (Set.range f.base)) : IsClosedImmersion f := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean index 4d5503fe1dc7f..a2505954c2c6e 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Preimmersion.lean @@ -4,8 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.Morphisms.UnderlyingMap -import Mathlib.RingTheory.RingHom.Surjective -import Mathlib.RingTheory.SurjectiveOnStalks +import Mathlib.AlgebraicGeometry.Morphisms.SurjectiveOnStalks /-! @@ -33,9 +32,8 @@ namespace AlgebraicGeometry /-- A morphism of schemes `f : X ⟶ Y` is a preimmersion if the underlying map of topological spaces is an embedding and the induced morphisms of stalks are all surjective. -/ @[mk_iff] -class IsPreimmersion {X Y : Scheme} (f : X ⟶ Y) : Prop where +class IsPreimmersion {X Y : Scheme} (f : X ⟶ Y) extends SurjectiveOnStalks f : Prop where base_embedding : IsEmbedding f.base - surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) lemma Scheme.Hom.isEmbedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : IsEmbedding f.base := IsPreimmersion.base_embedding @@ -43,12 +41,8 @@ lemma Scheme.Hom.isEmbedding {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] : I @[deprecated (since := "2024-10-26")] alias Scheme.Hom.embedding := Scheme.Hom.isEmbedding -lemma Scheme.Hom.stalkMap_surjective {X Y : Scheme} (f : Hom X Y) [IsPreimmersion f] (x) : - Function.Surjective (f.stalkMap x) := - IsPreimmersion.surj_on_stalks x - lemma isPreimmersion_eq_inf : - @IsPreimmersion = topologically IsEmbedding ⊓ stalkwise (Function.Surjective ·) := by + @IsPreimmersion = (@SurjectiveOnStalks ⊓ topologically IsEmbedding : MorphismProperty _) := by ext rw [isPreimmersion_iff] rfl @@ -69,10 +63,7 @@ instance (priority := 900) {X Y : Scheme} (f : X ⟶ Y) [IsOpenImmersion f] : Is instance : MorphismProperty.IsMultiplicative @IsPreimmersion where id_mem _ := inferInstance - comp_mem {X Y Z} f g hf hg := by - refine ⟨hg.base_embedding.comp hf.base_embedding, fun x ↦ ?_⟩ - rw [Scheme.stalkMap_comp] - exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) + comp_mem f g _ _ := ⟨g.isEmbedding.comp f.isEmbedding⟩ instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [IsPreimmersion f] [IsPreimmersion g] : IsPreimmersion (f ≫ g) := @@ -104,14 +95,8 @@ lemma Spec_map_iff {R S : CommRingCat.{u}} (f : R ⟶ S) : haveI : (RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).RespectsIso := by rw [← RingHom.toMorphismProperty_respectsIso_iff] exact RingHom.surjective_respectsIso - refine ⟨fun ⟨h₁, h₂⟩ ↦ ⟨h₁, ?_⟩, fun ⟨h₁, h₂⟩ ↦ ⟨h₁, fun (x : PrimeSpectrum S) ↦ ?_⟩⟩ - · intro p hp - let e := Scheme.arrowStalkMapSpecIso f ⟨p, hp⟩ - apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mp - exact h₂ ⟨p, hp⟩ - · let e := Scheme.arrowStalkMapSpecIso f x - apply ((RingHom.toMorphismProperty <| fun f ↦ Function.Surjective f).arrow_mk_iso_iff e).mpr - exact h₂ x.asIdeal x.isPrime + rw [← HasRingHomProperty.Spec_iff (P := @SurjectiveOnStalks), isPreimmersion_iff, and_comm] + rfl lemma mk_Spec_map {R S : CommRingCat.{u}} {f : R ⟶ S} (h₁ : IsEmbedding (PrimeSpectrum.comap f)) (h₂ : f.SurjectiveOnStalks) : @@ -125,6 +110,19 @@ lemma of_isLocalization {R S : Type u} [CommRing R] (M : Submonoid R) [CommRing (PrimeSpectrum.localization_comap_isEmbedding (R := R) S M) (RingHom.surjectiveOnStalks_of_isLocalization (M := M) S) +open Limits MorphismProperty in +instance : IsStableUnderBaseChange @IsPreimmersion := by + refine .mk' fun X Y Z f g _ _ ↦ ?_ + have := pullback_fst (P := @SurjectiveOnStalks) f g inferInstance + constructor + let L (x : (pullback f g : _)) : { x : X × Y | f.base x.1 = g.base x.2 } := + ⟨⟨(pullback.fst f g).base x, (pullback.snd f g).base x⟩, + by simp only [Set.mem_setOf, ← Scheme.comp_base_apply, pullback.condition]⟩ + have : IsEmbedding L := IsEmbedding.of_comp (by fun_prop) continuous_subtype_val + (SurjectiveOnStalks.isEmbedding_pullback f g) + exact IsEmbedding.subtypeVal.comp ((TopCat.pullbackHomeoPreimage _ f.continuous _ + g.isEmbedding).isEmbedding.comp this) + end IsPreimmersion end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean new file mode 100644 index 0000000000000..548f463af6395 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -0,0 +1,186 @@ +/- +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.RingHomProperties +import Mathlib.AlgebraicGeometry.PrimeSpectrum.TensorProduct +import Mathlib.Topology.LocalAtTarget + +/-! +# Morphisms surjective on stalks + +We define the classe of morphisms between schemes that are surjective on stalks. +We show that this class is stable under composition and base change. + +We also show that (`AlgebraicGeometry.SurjectiveOnStalks.isEmbedding_pullback`) +if `Y ⟶ S` is surjective on stalks, then for every `X ⟶ S`, `X ×ₛ Y` is a subset of +`X × Y` (cartesian product as topological spaces) with the induced topology. +-/ + +open CategoryTheory CategoryTheory.Limits Topology + +namespace AlgebraicGeometry + +universe u + +variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) + +/-- The class of morphisms `f : X ⟶ Y` between schemes such that +`𝒪_{Y, f x} ⟶ 𝒪_{X, x}` is surjective for all `x : X`. -/ +@[mk_iff] +class SurjectiveOnStalks : Prop where + surj_on_stalks : ∀ x, Function.Surjective (f.stalkMap x) + +theorem Scheme.Hom.stalkMap_surjective (f : X.Hom Y) [SurjectiveOnStalks f] (x) : + Function.Surjective (f.stalkMap x) := + SurjectiveOnStalks.surj_on_stalks x + +namespace SurjectiveOnStalks + +instance (priority := 900) [IsOpenImmersion f] : SurjectiveOnStalks f := + ⟨fun _ ↦ (ConcreteCategory.bijective_of_isIso _).2⟩ + +instance : MorphismProperty.IsMultiplicative @SurjectiveOnStalks where + id_mem _ := inferInstance + comp_mem {X Y Z} f g hf hg := by + refine ⟨fun x ↦ ?_⟩ + rw [Scheme.stalkMap_comp] + exact (hf.surj_on_stalks x).comp (hg.surj_on_stalks (f.base x)) + +instance comp {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) [SurjectiveOnStalks f] + [SurjectiveOnStalks g] : SurjectiveOnStalks (f ≫ g) := + MorphismProperty.IsStableUnderComposition.comp_mem f g inferInstance inferInstance + +lemma eq_stalkwise : + @SurjectiveOnStalks = stalkwise (Function.Surjective ·) := by + ext; exact surjectiveOnStalks_iff _ + +instance : IsLocalAtTarget @SurjectiveOnStalks := + eq_stalkwise ▸ stalkwiseIsLocalAtTarget_of_respectsIso RingHom.surjective_respectsIso + +instance : IsLocalAtSource @SurjectiveOnStalks := + eq_stalkwise ▸ stalkwise_isLocalAtSource_of_respectsIso RingHom.surjective_respectsIso + +lemma Spec_iff {R S : CommRingCat.{u}} {φ : R ⟶ S} : + SurjectiveOnStalks (Spec.map φ) ↔ RingHom.SurjectiveOnStalks φ := by + rw [eq_stalkwise, stalkwise_Spec_map_iff RingHom.surjective_respectsIso, + RingHom.SurjectiveOnStalks] + +instance : HasRingHomProperty @SurjectiveOnStalks RingHom.SurjectiveOnStalks := + eq_stalkwise ▸ .stalkwise RingHom.surjective_respectsIso + +variable {f} in +lemma iff_of_isAffine [IsAffine X] [IsAffine Y] : + SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤) := by + rw [← Spec_iff, MorphismProperty.arrow_mk_iso_iff @SurjectiveOnStalks (arrowIsoSpecΓOfIsAffine f)] + rfl + +theorem of_comp [SurjectiveOnStalks (f ≫ g)] : SurjectiveOnStalks f := by + refine ⟨fun x ↦ ?_⟩ + have := (f ≫ g).stalkMap_surjective x + rw [Scheme.stalkMap_comp] at this + exact Function.Surjective.of_comp this + +instance stableUnderBaseChange : + MorphismProperty.IsStableUnderBaseChange @SurjectiveOnStalks := by + apply HasRingHomProperty.isStableUnderBaseChange + apply RingHom.IsStableUnderBaseChange.mk + · exact (HasRingHomProperty.isLocal_ringHomProperty @SurjectiveOnStalks).respectsIso + intros R S T _ _ _ _ _ H + exact H.baseChange + +/-- If `Y ⟶ S` is surjective on stalks, then for every `X ⟶ S`, `X ×ₛ Y` is a subset of +`X × Y` (cartesian product as topological spaces) with the induced topology. -/ +lemma isEmbedding_pullback {X Y S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) [SurjectiveOnStalks g] : + IsEmbedding (fun x ↦ ((pullback.fst f g).base x, (pullback.snd f g).base x)) := by + let L := (fun x ↦ ((pullback.fst f g).base x, (pullback.snd f g).base x)) + have H : ∀ R A B (f' : Spec A ⟶ Spec R) (g' : Spec B ⟶ Spec R) (iX : Spec A ⟶ X) + (iY : Spec B ⟶ Y) (iS : Spec R ⟶ S) (e₁ e₂), IsOpenImmersion iX → IsOpenImmersion iY → + IsOpenImmersion iS → IsEmbedding (L ∘ (pullback.map f' g' f g iX iY iS e₁ e₂).base) := by + intro R A B f' g' iX iY iS e₁ e₂ _ _ _ + have H : SurjectiveOnStalks g' := + have : SurjectiveOnStalks (g' ≫ iS) := e₂ ▸ inferInstance + .of_comp _ iS + obtain ⟨φ, rfl⟩ : ∃ φ, Spec.map φ = f' := ⟨_, Spec.map_preimage _⟩ + obtain ⟨ψ, rfl⟩ : ∃ ψ, Spec.map ψ = g' := ⟨_, Spec.map_preimage _⟩ + letI := φ.toAlgebra + letI := ψ.toAlgebra + rw [HasRingHomProperty.Spec_iff (P := @SurjectiveOnStalks)] at H + convert ((iX.isOpenEmbedding.prodMap iY.isOpenEmbedding).isEmbedding.comp + (PrimeSpectrum.isEmbedding_tensorProductTo_of_surjectiveOnStalks R A B H)).comp + (Scheme.homeoOfIso (pullbackSpecIso R A B)).isEmbedding + ext1 x + obtain ⟨x, rfl⟩ := (Scheme.homeoOfIso (pullbackSpecIso R A B).symm).surjective x + simp only [Scheme.homeoOfIso_apply, Function.comp_apply] + ext + · simp only [← Scheme.comp_base_apply, pullback.lift_fst, Iso.symm_hom, Iso.inv_hom_id] + erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_fst_assoc] + rfl + · simp only [← Scheme.comp_base_apply, pullback.lift_snd, Iso.symm_hom, Iso.inv_hom_id] + erw [← Scheme.comp_base_apply, pullbackSpecIso_inv_snd_assoc] + rfl + let 𝒰 := S.affineOpenCover.openCover + let 𝒱 (i) := ((𝒰.pullbackCover f).obj i).affineOpenCover.openCover + let 𝒲 (i) := ((𝒰.pullbackCover g).obj i).affineOpenCover.openCover + let U (ijk : Σ i, (𝒱 i).J × (𝒲 i).J) : TopologicalSpace.Opens (X.carrier × Y) := + ⟨{ P | P.1 ∈ ((𝒱 ijk.1).map ijk.2.1 ≫ (𝒰.pullbackCover f).map ijk.1).opensRange ∧ + P.2 ∈ ((𝒲 ijk.1).map ijk.2.2 ≫ (𝒰.pullbackCover g).map ijk.1).opensRange }, + (continuous_fst.1 _ ((𝒱 ijk.1).map ijk.2.1 ≫ + (𝒰.pullbackCover f).map ijk.1).opensRange.2).inter (continuous_snd.1 _ + ((𝒲 ijk.1).map ijk.2.2 ≫ (𝒰.pullbackCover g).map ijk.1).opensRange.2)⟩ + have : Set.range L ⊆ (iSup U : _) := by + simp only [Scheme.Cover.pullbackCover_J, Scheme.Cover.pullbackCover_obj, Set.range_subset_iff] + intro z + simp only [SetLike.mem_coe, TopologicalSpace.Opens.mem_iSup, Sigma.exists, Prod.exists] + obtain ⟨is, s, hsx⟩ := 𝒰.exists_eq (f.base ((pullback.fst f g).base z)) + have hsy : (𝒰.map is).base s = g.base ((pullback.snd f g).base z) := by + rwa [← Scheme.comp_base_apply, ← pullback.condition, Scheme.comp_base_apply] + obtain ⟨x : (𝒰.pullbackCover f).obj is, hx⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hsx.symm + obtain ⟨y : (𝒰.pullbackCover g).obj is, hy⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hsy.symm + obtain ⟨ix, x, rfl⟩ := (𝒱 is).exists_eq x + obtain ⟨iy, y, rfl⟩ := (𝒲 is).exists_eq y + refine ⟨is, ix, iy, ⟨x, hx⟩, ⟨y, hy⟩⟩ + let 𝓤 := (Scheme.Pullback.openCoverOfBase 𝒰 f g).bind + (fun i ↦ Scheme.Pullback.openCoverOfLeftRight (𝒱 i) (𝒲 i) _ _) + refine isEmbedding_of_iSup_eq_top_of_preimage_subset_range _ ?_ U this _ (fun i ↦ (𝓤.map i).base) + (fun i ↦ (𝓤.map i).continuous) ?_ ?_ + · fun_prop + · rintro i x ⟨⟨x₁, hx₁⟩, ⟨x₂, hx₂⟩⟩ + obtain ⟨x₁', hx₁'⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hx₁.symm + obtain ⟨x₂', hx₂'⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ hx₂.symm + obtain ⟨z, hz⟩ := + Scheme.IsJointlySurjectivePreserving.exists_preimage_fst_triplet_of_prop + (P := @IsOpenImmersion) inferInstance _ _ (hx₁'.trans hx₂'.symm) + refine ⟨(pullbackFstFstIso _ _ _ _ _ _ (𝒰.map i.1) ?_ ?_).hom.base z, ?_⟩ + · simp [pullback.condition] + · simp [pullback.condition] + · dsimp only + rw [← hx₁', ← hz, ← Scheme.comp_base_apply] + erw [← Scheme.comp_base_apply] + congr 4 + apply pullback.hom_ext <;> simp [𝓤, ← pullback.condition, ← pullback.condition_assoc] + · intro i + have := H (S.affineOpenCover.obj i.1) (((𝒰.pullbackCover f).obj i.1).affineOpenCover.obj i.2.1) + (((𝒰.pullbackCover g).obj i.1).affineOpenCover.obj i.2.2) + ((𝒱 i.1).map i.2.1 ≫ 𝒰.pullbackHom f i.1) + ((𝒲 i.1).map i.2.2 ≫ 𝒰.pullbackHom g i.1) + ((𝒱 i.1).map i.2.1 ≫ (𝒰.pullbackCover f).map i.1) + ((𝒲 i.1).map i.2.2 ≫ (𝒰.pullbackCover g).map i.1) + (𝒰.map i.1) (by simp [pullback.condition]) (by simp [pullback.condition]) + inferInstance inferInstance inferInstance + convert this using 6 + apply pullback.hom_ext <;> + simp [𝓤, ← pullback.condition, ← pullback.condition_assoc, Scheme.Cover.pullbackHom] + +end SurjectiveOnStalks + +end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index f6caa6dd8c1f7..6decb4e3cb660 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -214,6 +214,14 @@ def forgetToTop : Scheme ⥤ TopCat := noncomputable def homeoOfIso {X Y : Scheme.{u}} (e : X ≅ Y) : X ≃ₜ Y := TopCat.homeoOfIso (forgetToTop.mapIso e) +@[simp] +lemma homeoOfIso_symm {X Y : Scheme} (e : X ≅ Y) : + (homeoOfIso e).symm = homeoOfIso e.symm := rfl + +@[simp] +lemma homeoOfIso_apply {X Y : Scheme} (e : X ≅ Y) (x : X) : + homeoOfIso e x = e.hom.base x := rfl + alias _root_.CategoryTheory.Iso.schemeIsoToHomeo := homeoOfIso /-- An isomorphism of schemes induces a homeomorphism of the underlying topological spaces. -/ @@ -221,6 +229,10 @@ noncomputable def Hom.homeomorph {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := S X ≃ₜ Y := (asIso f).schemeIsoToHomeo +@[simp] +lemma Hom.homeomorph_apply {X Y : Scheme.{u}} (f : X.Hom Y) [IsIso (C := Scheme) f] (x) : + f.homeomorph x = f.base x := rfl + -- Porting note: Lean seems not able to find this coercion any more instance hasCoeToTopCat : CoeOut Scheme TopCat where coe X := X.carrier diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 9d801889f8176..9ab72631191fb 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -185,6 +185,47 @@ theorem isEmbedding_iff_of_iSup_eq_top (h : Continuous f) : convert congr_arg SetLike.coe hU simp +omit hU in +/-- +Given a continuous map `f : X → Y` between topological spaces. +Suppose we have an open cover `V i` of the range of `f`, and an open cover `U i` of `X` that is +coarser than the pullback of `V` under `f`. +To check that `f` is an embedding it suffices to check that `U i → Y` is an embedding for all `i`. +-/ +theorem isEmbedding_of_iSup_eq_top_of_preimage_subset_range + {X Y} [TopologicalSpace X] [TopologicalSpace Y] + (f : X → Y) (h : Continuous f) {ι : Type*} + (U : ι → Opens Y) (hU : Set.range f ⊆ (iSup U : _)) + (V : ι → Type*) [∀ i, TopologicalSpace (V i)] + (iV : ∀ i, V i → X) (hiV : ∀ i, Continuous (iV i)) (hV : ∀ i, f ⁻¹' U i ⊆ Set.range (iV i)) + (hV' : ∀ i, IsEmbedding (f ∘ iV i)) : IsEmbedding f := by + wlog hU' : iSup U = ⊤ + · let f₀ : X → Set.range f := fun x ↦ ⟨f x, ⟨x, rfl⟩⟩ + suffices IsEmbedding f₀ from IsEmbedding.subtypeVal.comp this + have hU'' : (⨆ i, (U i).comap ⟨Subtype.val, continuous_subtype_val⟩ : + Opens (Set.range f)) = ⊤ := by + rw [← top_le_iff] + simpa [Set.range_subset_iff, SetLike.le_def] using hU + refine this _ ?_ _ ?_ V iV hiV ?_ ?_ hU'' + · fun_prop + · rw [hU'']; simp + · exact hV + · exact fun i ↦ IsEmbedding.of_comp (by fun_prop) continuous_subtype_val (hV' i) + rw [isEmbedding_iff_of_iSup_eq_top hU' h] + intro i + let f' := (Subtype.val ∘ (f ⁻¹' U i).restrictPreimage (iV i)) + have : IsEmbedding f' := + IsEmbedding.subtypeVal.comp ((IsEmbedding.of_comp (hiV i) h (hV' _)).restrictPreimage _) + have hf' : Set.range f' = f ⁻¹' U i := by + simpa [f', Set.range_comp, Set.range_restrictPreimage] using hV i + let e := (Homeomorph.ofIsEmbedding _ this).trans (Homeomorph.setCongr hf') + refine IsEmbedding.of_comp (by fun_prop) continuous_subtype_val ?_ + convert ((hV' i).comp IsEmbedding.subtypeVal).comp e.symm.isEmbedding + ext x + obtain ⟨x, rfl⟩ := e.surjective x + simp + rfl + @[deprecated (since := "2024-10-26")] alias embedding_iff_embedding_of_iSup_eq_top := isEmbedding_iff_of_iSup_eq_top From 6343082c7dd00a42b18e306d26a4fa513cd044d3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Violeta=20Hern=C3=A1ndez?= Date: Sat, 23 Nov 2024 14:58:16 +0000 Subject: [PATCH 085/250] feat: Sedrakyan's lemma (#19311) This is a specialization of the Cauchy-Schwarz inequality which is often useful in math olympiad problems. In order to prove it in general ordered semifields, we have to show slightly more general forms of the AM-GM and Cauchy-Schwarz inequalities, which indirectly work with square roots. Co-authored-by: Alex Brodbelt <64128135+AlexBrodbelt@users.noreply.github.com> Co-authored-by: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> --- .../Order/BigOperators/Ring/Finset.lean | 102 +++++++++++------- .../Algebra/Order/Ring/Unbundled/Basic.lean | 12 ++- 2 files changed, 75 insertions(+), 39 deletions(-) diff --git a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean index b51accd2d82ad..1c522e8948d04 100644 --- a/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean +++ b/Mathlib/Algebra/Order/BigOperators/Ring/Finset.lean @@ -14,9 +14,12 @@ import Mathlib.Tactic.Ring This file contains the results concerning the interaction of finset big operators with ordered rings. + +In particular, this file contains the standard form of the Cauchy-Schwarz inequality, as well as +some of its immediate consequences. -/ -variable {ι R : Type*} +variable {ι R S : Type*} namespace Finset @@ -120,41 +123,10 @@ lemma prod_add_prod_le {i : ι} {f g h : ι → R} (hi : i ∈ s) (h2i : g i + h end OrderedCommSemiring -section LinearOrderedCommSemiring -variable [LinearOrderedCommSemiring R] [ExistsAddOfLE R] - -/-- **Cauchy-Schwarz inequality** for finsets. -/ -lemma sum_mul_sq_le_sq_mul_sq (s : Finset ι) (f g : ι → R) : - (∑ i ∈ s, f i * g i) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) * ∑ i ∈ s, g i ^ 2 := by - nontriviality R - obtain h' | h' := (sum_nonneg fun _ _ ↦ sq_nonneg <| g _).eq_or_lt - · have h'' : ∀ i ∈ s, g i = 0 := fun i hi ↦ by - simpa using (sum_eq_zero_iff_of_nonneg fun i _ ↦ sq_nonneg (g i)).1 h'.symm i hi - rw [← h', sum_congr rfl (show ∀ i ∈ s, f i * g i = 0 from fun i hi ↦ by simp [h'' i hi])] - simp - refine le_of_mul_le_mul_of_pos_left - (le_of_add_le_add_left (a := (∑ i ∈ s, g i ^ 2) * (∑ j ∈ s, f j * g j) ^ 2) ?_) h' - calc - _ = ∑ i ∈ s, 2 * (f i * ∑ j ∈ s, g j ^ 2) * (g i * ∑ j ∈ s, f j * g j) := by - simp_rw [mul_assoc (2 : R), mul_mul_mul_comm, ← mul_sum, ← sum_mul]; ring - _ ≤ ∑ i ∈ s, ((f i * ∑ j ∈ s, g j ^ 2) ^ 2 + (g i * ∑ j ∈ s, f j * g j) ^ 2) := - 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 +theorem sum_mul_self_eq_zero_iff [LinearOrderedSemiring R] [ExistsAddOfLE R] (s : Finset ι) + (f : ι → R) : ∑ i ∈ s, f i * f i = 0 ↔ ∀ i ∈ s, f i = 0 := by + rw [sum_eq_zero_iff_of_nonneg fun _ _ ↦ mul_self_nonneg _] + simp lemma abs_prod [LinearOrderedCommRing R] (s : Finset ι) (f : ι → R) : |∏ x ∈ s, f x| = ∏ x ∈ s, |f x| := @@ -184,11 +156,63 @@ lemma prod_add_prod_le' (hi : i ∈ s) (h2i : g i + h i ≤ f i) (hgf : ∀ j assumption end CanonicallyOrderedCommSemiring + +/-! ### Named inequalities -/ + +/-- **Cauchy-Schwarz inequality** for finsets. + +This is written in terms of sequences `f`, `g`, and `r`, where `r` is a stand-in for +`√(f i * g i)`. See `sum_mul_sq_le_sq_mul_sq` for the more usual form in terms of squared +sequences. -/ +lemma sum_sq_le_sum_mul_sum_of_sq_eq_mul [LinearOrderedCommSemiring R] [ExistsAddOfLE R] + (s : Finset ι) {r f g : ι → R} (hf : ∀ i ∈ s, 0 ≤ f i) (hg : ∀ i ∈ s, 0 ≤ g i) + (ht : ∀ i ∈ s, r i ^ 2 = f i * g i) : (∑ i ∈ s, r i) ^ 2 ≤ (∑ i ∈ s, f i) * ∑ i ∈ s, g i := by + obtain h | h := (sum_nonneg hg).eq_or_gt + · have ht' : ∑ i ∈ s, r i = 0 := sum_eq_zero fun i hi ↦ by + simpa [(sum_eq_zero_iff_of_nonneg hg).1 h i hi] using ht i hi + rw [h, ht'] + simp + · refine le_of_mul_le_mul_of_pos_left + (le_of_add_le_add_left (a := (∑ i ∈ s, g i) * (∑ i ∈ s, r i) ^ 2) ?_) h + calc + _ = ∑ i ∈ s, 2 * r i * (∑ j ∈ s, g j) * (∑ j ∈ s, r j) := by + simp_rw [mul_assoc, ← mul_sum, ← sum_mul]; ring + _ ≤ ∑ i ∈ s, (f i * (∑ j ∈ s, g j) ^ 2 + g i * (∑ j ∈ s, r j) ^ 2) := by + gcongr with i hi + have ht : (r i * (∑ j ∈ s, g j) * (∑ j ∈ s, r j)) ^ 2 = + (f i * (∑ j ∈ s, g j) ^ 2) * (g i * (∑ j ∈ s, r j) ^ 2) := by + conv_rhs => rw [mul_mul_mul_comm, ← ht i hi] + ring + refine le_of_eq_of_le ?_ (two_mul_le_add_of_sq_eq_mul + (mul_nonneg (hf i hi) (sq_nonneg _)) (mul_nonneg (hg i hi) (sq_nonneg _)) ht) + repeat rw [mul_assoc] + _ = _ := by simp_rw [sum_add_distrib, ← sum_mul]; ring + +/-- **Cauchy-Schwarz inequality** for finsets, squared version. -/ +lemma sum_mul_sq_le_sq_mul_sq [LinearOrderedCommSemiring R] [ExistsAddOfLE R] (s : Finset ι) + (f g : ι → R) : (∑ i ∈ s, f i * g i) ^ 2 ≤ (∑ i ∈ s, f i ^ 2) * ∑ i ∈ s, g i ^ 2 := + sum_sq_le_sum_mul_sum_of_sq_eq_mul s + (fun _ _ ↦ sq_nonneg _) (fun _ _ ↦ sq_nonneg _) (fun _ _ ↦ mul_pow ..) + +/-- **Sedrakyan's lemma**, aka **Titu's lemma** or **Engel's form**. + +This is a specialization of the Cauchy-Schwarz inequality with the sequences `f n / √(g n)` and +`√(g n)`, though here it is proven without relying on square roots. -/ +theorem sq_sum_div_le_sum_sq_div [LinearOrderedSemifield R] [ExistsAddOfLE R] (s : Finset ι) + (f : ι → R) {g : ι → R} (hg : ∀ i ∈ s, 0 < g i) : + (∑ i ∈ s, f i) ^ 2 / ∑ i ∈ s, g i ≤ ∑ i ∈ s, f i ^ 2 / g i := by + have hg' : ∀ i ∈ s, 0 ≤ g i := fun i hi ↦ (hg i hi).le + have H : ∀ i ∈ s, 0 ≤ f i ^ 2 / g i := fun i hi ↦ div_nonneg (sq_nonneg _) (hg' i hi) + refine div_le_of_le_mul₀ (sum_nonneg hg') (sum_nonneg H) + (sum_sq_le_sum_mul_sum_of_sq_eq_mul _ H hg' fun i hi ↦ ?_) + rw [div_mul_cancel₀] + exact (hg i hi).ne' + end Finset -section AbsoluteValue +/-! ### Absolute values -/ -variable {S : Type*} +section AbsoluteValue lemma AbsoluteValue.sum_le [Semiring R] [OrderedSemiring S] (abv : AbsoluteValue R S) (s : Finset ι) (f : ι → R) : abv (∑ i ∈ s, f i) ≤ ∑ i ∈ s, abv (f i) := @@ -212,6 +236,8 @@ lemma IsAbsoluteValue.map_prod [CommSemiring R] [Nontrivial R] [LinearOrderedCom end AbsoluteValue +/-! ### Positivity extension -/ + namespace Mathlib.Meta.Positivity open Qq Lean Meta Finset diff --git a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean index 06cc0315f08b0..99bd1f2cfd31c 100644 --- a/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean +++ b/Mathlib/Algebra/Order/Ring/Unbundled/Basic.lean @@ -764,7 +764,7 @@ alias two_mul_le_add_pow_two := two_mul_le_add_sq /-- Binary, squared, and division-free **arithmetic mean-geometric mean inequality** (aka AM-GM inequality) for linearly ordered commutative semirings. -/ lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] - [ContravariantClass α α (· + ·) (· ≤ ·)] [CovariantClass α α (· + ·) (· ≤ ·)] + [AddLeftReflectLE α] [AddLeftMono α] (a b : α) : 4 * a * b ≤ (a + b) ^ 2 := by calc 4 * a * b _ = 2 * a * b + 2 * a * b := by rw [mul_assoc, two_add_two_eq_four.symm, add_mul, mul_assoc] @@ -774,6 +774,16 @@ lemma four_mul_le_sq_add [ExistsAddOfLE α] [MulPosStrictMono α] alias four_mul_le_pow_two_add := four_mul_le_sq_add +/-- Binary and division-free **arithmetic mean-geometric mean inequality** +(aka AM-GM inequality) for linearly ordered commutative semirings. -/ +lemma two_mul_le_add_of_sq_eq_mul [ExistsAddOfLE α] [MulPosStrictMono α] [PosMulStrictMono α] + [AddLeftReflectLE α] [AddLeftMono α] {a b r : α} + (ha : 0 ≤ a) (hb : 0 ≤ b) (ht : r ^ 2 = a * b) : 2 * r ≤ a + b := by + apply nonneg_le_nonneg_of_sq_le_sq (Left.add_nonneg ha hb) + conv_rhs => rw [← pow_two] + convert four_mul_le_sq_add a b using 1 + rw [mul_mul_mul_comm, two_mul, two_add_two_eq_four, ← pow_two, ht, mul_assoc] + end LinearOrderedCommSemiring section LinearOrderedRing From cc6b7ff0c02212e54b56437933156ee1cdf452fe Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 15:57:51 +0000 Subject: [PATCH 086/250] =?UTF-8?q?feat:=20`s=E2=81=BB=C2=B9.encard=20=3D?= =?UTF-8?q?=20s.encard`=20(#19400)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From Kneser (LeanCamCombi) --- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 80a073b792385..06a3718af24d0 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,6 +3,7 @@ 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.Data.Set.Card import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite @@ -49,6 +50,12 @@ lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by @[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv +@[to_additive (attr := simp)] +lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, PartENat.card] + +@[to_additive (attr := simp)] +lemma ncard_inv (s : Set G) : s⁻¹.ncard = s.ncard := by simp [ncard] + end InvolutiveInv section DivInvMonoid @@ -82,5 +89,12 @@ lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s @[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_smul_set := Cardinal.mk_smul_set +@[to_additive (attr := simp)] +lemma encard_smul_set (a : G) (s : Set α) : (a • s).encard = s.encard := by + simp [encard, PartENat.card] + +@[to_additive (attr := simp)] +lemma ncard_smul_set (a : G) (s : Set α) : (a • s).ncard = s.ncard := by simp [ncard] + end Group end Set From 571c2829083a7dcf8934dff094f89869b3f6244f Mon Sep 17 00:00:00 2001 From: damiano Date: Sat, 23 Nov 2024 19:35:47 +0000 Subject: [PATCH 087/250] chore: remove unused `Category` assumption (#19416) The `Category` assumption is unused, as `Discrete` subsumes it. --- Mathlib/CategoryTheory/Groupoid/Discrete.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Groupoid/Discrete.lean b/Mathlib/CategoryTheory/Groupoid/Discrete.lean index ffface4377069..3ce9a350831bf 100644 --- a/Mathlib/CategoryTheory/Groupoid/Discrete.lean +++ b/Mathlib/CategoryTheory/Groupoid/Discrete.lean @@ -12,7 +12,7 @@ import Mathlib.CategoryTheory.DiscreteCategory namespace CategoryTheory -variable {C : Type*} [Category C] +variable {C : Type*} instance : Groupoid (Discrete C) := { inv := fun h ↦ ⟨⟨h.1.1.symm⟩⟩ } From fee45b61fbab23192de1eb6f8f536e53722254f1 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sat, 23 Nov 2024 19:55:50 +0000 Subject: [PATCH 088/250] feat(AlgebraicGeometry): the domain of definition of a rational map (#18803) --- Mathlib/AlgebraicGeometry/AffineSpace.lean | 2 +- Mathlib/AlgebraicGeometry/Limits.lean | 6 +- .../Morphisms/Separated.lean | 10 + Mathlib/AlgebraicGeometry/Over.lean | 4 + Mathlib/AlgebraicGeometry/RationalMap.lean | 220 +++++++++++++++++- Mathlib/AlgebraicGeometry/Restrict.lean | 21 +- Mathlib/CategoryTheory/Comma/OverClass.lean | 20 +- 7 files changed, 273 insertions(+), 10 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 6374a6564ad9f..8e43ca277f77f 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -332,7 +332,7 @@ def mapSpecMap {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : def reindex {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : 𝔸(n; S) ⟶ 𝔸(m; S) := homOfVector (𝔸(n; S) ↘ S) (coord S ∘ i) -@[reassoc (attr := simp)] +@[simp, reassoc] lemma reindex_over {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : reindex i S ≫ 𝔸(m; S) ↘ S = 𝔸(n; S) ↘ S := pullback.lift_fst _ _ _ diff --git a/Mathlib/AlgebraicGeometry/Limits.lean b/Mathlib/AlgebraicGeometry/Limits.lean index 6c634b3aadaac..b4f7a0106ddca 100644 --- a/Mathlib/AlgebraicGeometry/Limits.lean +++ b/Mathlib/AlgebraicGeometry/Limits.lean @@ -48,7 +48,11 @@ instance : HasFiniteLimits Scheme := hasFiniteLimits_of_hasTerminal_and_pullbacks instance (X : Scheme.{u}) : X.Over (⊤_ _) := ⟨terminal.from _⟩ -instance {X Y : Scheme.{u}} (f : X ⟶ Y) : f.IsOver (⊤_ _) := ⟨terminal.hom_ext _ _⟩ +instance {X Y : Scheme.{u}} [X.Over (⊤_ Scheme)] [Y.Over (⊤_ Scheme)] (f : X ⟶ Y) : + @Scheme.Hom.IsOver _ _ f (⊤_ Scheme) ‹_› ‹_› := ⟨Subsingleton.elim _ _⟩ + +instance {X : Scheme} : Subsingleton (X.Over (⊤_ Scheme)) := + ⟨fun ⟨a⟩ ⟨b⟩ ↦ by simp [Subsingleton.elim a b]⟩ section Initial diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean index a97b12f9fcc9c..c03365e68f13f 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Separated.lean @@ -240,6 +240,16 @@ lemma ext_of_isDominant_of_isSeparated [IsReduced X] {f g : X ⟶ Y} rw [← cancel_epi (equalizer.ι f' g').left] exact congr($(equalizer.condition f' g').left) +variable (S) in +/-- +Suppose `X` is a reduced `S`-scheme and `Y` is a separated `S`-scheme. +For any `S`-morphisms `f g : X ⟶ Y`, `f = g` if `ι ≫ f = ι ≫ g` for some dominant `ι`. +-/ +lemma ext_of_isDominant_of_isSeparated' [X.Over S] [Y.Over S] [IsReduced X] [IsSeparated (Y ↘ S)] + {f g : X ⟶ Y} [f.IsOver S] [g.IsOver S] {W} (ι : W ⟶ X) [IsDominant ι] + (hU : ι ≫ f = ι ≫ g) : f = g := + ext_of_isDominant_of_isSeparated (Y ↘ S) (by simp) ι hU + namespace Scheme /-- A scheme `X` is separated if it is separated over `⊤_ Scheme`. -/ diff --git a/Mathlib/AlgebraicGeometry/Over.lean b/Mathlib/AlgebraicGeometry/Over.lean index c52efcbb93841..b0f3bcaef6d0e 100644 --- a/Mathlib/AlgebraicGeometry/Over.lean +++ b/Mathlib/AlgebraicGeometry/Over.lean @@ -41,6 +41,10 @@ abbrev CanonicallyOver := CanonicallyOverClass X S `f.IsOver S` is the typeclass asserting `f` commutes with the structure morphisms. -/ abbrev Hom.IsOver (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] := HomIsOver f S +@[simp] +lemma Hom.isOver_iff [X.Over S] [Y.Over S] {f : X ⟶ Y} : f.IsOver S ↔ f ≫ Y ↘ S = X ↘ S := + ⟨fun H ↦ H.1, fun h ↦ ⟨h⟩⟩ + /-! Also note the existence of `CategoryTheory.IsOverTower X Y S`. -/ end AlgebraicGeometry.Scheme diff --git a/Mathlib/AlgebraicGeometry/RationalMap.lean b/Mathlib/AlgebraicGeometry/RationalMap.lean index aca42607f4a77..e9894ddfbcd61 100644 --- a/Mathlib/AlgebraicGeometry/RationalMap.lean +++ b/Mathlib/AlgebraicGeometry/RationalMap.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.AlgebraicGeometry.SpreadingOut import Mathlib.AlgebraicGeometry.FunctionField +import Mathlib.AlgebraicGeometry.Morphisms.Separated /-! # Rational maps between schemes @@ -17,10 +18,12 @@ import Mathlib.AlgebraicGeometry.FunctionField Two partial maps are equivalent if they are equal on a dense open subscheme. * `AlgebraicGeometry.Scheme.RationalMap`: A rational map from `X` to `Y` (`X ⤏ Y`) is an equivalence class of partial maps. -* `AlgebraicGeometry.Scheme.RationalMap.equivFunctionField`: +* `AlgebraicGeometry.Scheme.RationalMap.equivFunctionFieldOver`: Given `S`-schemes `X` and `Y` such that `Y` is locally of finite type and `X` is integral, `S`-morphisms `Spec K(X) ⟶ Y` correspond bijectively to `S`-rational maps from `X` to `Y`. - +* `AlgebraicGeometry.Scheme.RationalMap.toPartialMap`: + If `X` is integral and `Y` is separated, then any `f : X ⤏ Y` can be realized as a partial + map on `f.domain`, the domain of definition of `f`. -/ universe u @@ -44,6 +47,11 @@ structure PartialMap (X Y : Scheme.{u}) where /-- The underlying morphism of a partial map. -/ hom : ↑domain ⟶ Y +variable (S) in +/-- A partial map is a `S`-map if the underlying morphism is. -/ +abbrev PartialMap.IsOver [X.Over S] [Y.Over S] (f : X.PartialMap Y) := + f.hom.IsOver S + namespace PartialMap lemma ext_iff (f g : X.PartialMap Y) : @@ -64,7 +72,7 @@ lemma ext (f g : X.PartialMap Y) (e : f.domain = g.domain) exact ⟨e, H⟩ /-- The restriction of a partial map to a smaller domain. -/ -@[simps hom, simps (config := .lemmasOnly) domain] +@[simps hom domain] noncomputable def restrict (f : X.PartialMap Y) (U : X.Opens) (hU : Dense (U : Set X)) (hU' : U ≤ f.domain) : X.PartialMap Y where @@ -93,6 +101,10 @@ lemma restrict_restrict_hom (f : X.PartialMap Y) ((f.restrict U hU hU').restrict V hV hV').hom = (f.restrict V hV (hV'.trans hU')).hom := by simp +instance [X.Over S] [Y.Over S] (f : X.PartialMap Y) [f.IsOver S] + (U : X.Opens) (hU : Dense (U : Set X)) (hU' : U ≤ f.domain) : + (f.restrict U hU hU').IsOver S where + /-- The composition of a partial map and a morphism on the right. -/ @[simps] def compHom (f : X.PartialMap Y) (g : Y ⟶ Z) : X.PartialMap Z where @@ -100,11 +112,24 @@ def compHom (f : X.PartialMap Y) (g : Y ⟶ Z) : X.PartialMap Z where dense_domain := f.dense_domain hom := f.hom ≫ g +instance [X.Over S] [Y.Over S] [Z.Over S] (f : X.PartialMap Y) (g : Y ⟶ Z) + [f.IsOver S] [g.IsOver S] : (f.compHom g).IsOver S where + /-- A scheme morphism as a partial map. -/ @[simps] def _root_.AlgebraicGeometry.Scheme.Hom.toPartialMap (f : X.Hom Y) : X.PartialMap Y := ⟨⊤, dense_univ, X.topIso.hom ≫ f⟩ +instance [X.Over S] [Y.Over S] (f : X ⟶ Y) [f.IsOver S] : f.toPartialMap.IsOver S where + +lemma isOver_iff [X.Over S] [Y.Over S] {f : X.PartialMap Y} : + f.IsOver S ↔ (f.compHom (Y ↘ S)).hom = f.domain.ι ≫ X ↘ S := by + simp + +lemma isOver_iff_eq_restrict [X.Over S] [Y.Over S] {f : X.PartialMap Y} : + f.IsOver S ↔ f.compHom (Y ↘ S) = (X ↘ S).toPartialMap.restrict _ f.dense_domain (by simp) := by + simp [isOver_iff, PartialMap.ext_iff] + /-- If `x` is in the domain of a partial map `f`, then `f` restricts to a map from `Spec 𝒪_x`. -/ noncomputable def fromSpecStalkOfMem (f : X.PartialMap Y) {x} (hx : x ∈ f.domain) : @@ -230,6 +255,58 @@ lemma equiv_of_fromSpecStalkOfMem_eq [IrreducibleSpace X] simpa only [fromSpecStalkOfMem, restrict_domain, Opens.fromSpecStalkOfMem, Spec.map_inv, restrict_hom, Category.assoc, IsIso.eq_inv_comp, IsIso.hom_inv_id_assoc] using H +instance (U : X.Opens) [IsReduced X] : IsReduced U := isReduced_of_isOpenImmersion U.ι + +lemma Opens.isDominant_ι {U : X.Opens} (hU : Dense (X := X) U) : IsDominant U.ι := + ⟨by simpa [DenseRange] using hU⟩ + +lemma Opens.isDominant_homOfLE {U V : X.Opens} (hU : Dense (X := X) U) (hU' : U ≤ V) : + IsDominant (X.homOfLE hU') := + have : IsDominant (X.homOfLE hU' ≫ Opens.ι _) := by simpa using Opens.isDominant_ι hU + IsDominant.of_comp_of_isOpenImmersion (g := Opens.ι _) _ + +/-- Two partial maps from reduced schemes to separated schemes are equivalent if and only if +they are equal on **any** open dense subset. -/ +lemma equiv_iff_of_isSeparated_of_le [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} [f.IsOver S] [g.IsOver S] + {W : X.Opens} (hW : Dense (X := X) W) (hWl : W ≤ f.domain) (hWr : W ≤ g.domain) : f.equiv g ↔ + (f.restrict W hW hWl).hom = (g.restrict W hW hWr).hom := by + refine ⟨fun ⟨V, hV, hVl, hVr, e⟩ ↦ ?_, fun e ↦ ⟨_, _, _, _, e⟩⟩ + have : IsDominant (X.homOfLE (inf_le_left : W ⊓ V ≤ W)) := + Opens.isDominant_homOfLE (hW.inter_of_isOpen_left hV W.2) _ + apply ext_of_isDominant_of_isSeparated' S (X.homOfLE (inf_le_left : W ⊓ V ≤ W)) + simpa using congr(X.homOfLE (inf_le_right : W ⊓ V ≤ V) ≫ $e) + +/-- Two partial maps from reduced schemes to separated schemes are equivalent if and only if +they are equal on the intersection of the domains. -/ +lemma equiv_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} + [f.IsOver S] [g.IsOver S] : f.equiv g ↔ + (f.restrict _ (f.2.inter_of_isOpen_left g.2 f.domain.2) inf_le_left).hom = + (g.restrict _ (f.2.inter_of_isOpen_left g.2 f.domain.2) inf_le_right).hom := + equiv_iff_of_isSeparated_of_le (S := S) _ _ _ + +/-- Two partial maps from reduced schemes to separated schemes with the same domain are equivalent +if and only if they are equal. -/ +lemma equiv_iff_of_domain_eq_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f g : X.PartialMap Y} (hfg : f.domain = g.domain) + [f.IsOver S] [g.IsOver S] : f.equiv g ↔ f = g := by + rw [equiv_iff_of_isSeparated_of_le (S := S) f.dense_domain le_rfl hfg.le] + obtain ⟨Uf, _, f⟩ := f + obtain ⟨Ug, _, g⟩ := g + obtain rfl : Uf = Ug := hfg + simp + +/-- A partial map from a reduced scheme to a separated scheme is equivalent to a morphism +if and only if it is equal to the restriction of the morphism. -/ +lemma equiv_toPartialMap_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [IsSeparated (Y ↘ S)] {f : X.PartialMap Y} {g : X ⟶ Y} + [f.IsOver S] [g.IsOver S] : f.equiv g.toPartialMap ↔ + f.hom = f.domain.ι ≫ g := by + rw [equiv_iff_of_isSeparated (S := S), ← cancel_epi (X.isoOfEq (inf_top_eq f.domain)).hom] + simp + rfl + end PartialMap /-- A rational map from `X` to `Y` (`X ⤏ Y`) is an equivalence class of partial maps, @@ -246,6 +323,11 @@ def PartialMap.toRationalMap (f : X.PartialMap Y) : X ⤏ Y := Quotient.mk _ f /-- A scheme morphism as a rational map. -/ abbrev Hom.toRationalMap (f : X.Hom Y) : X ⤏ Y := f.toPartialMap.toRationalMap +variable (S) in +/-- A rational map is a `S`-map if some partial map in the equivalence class is a `S`-map. -/ +class RationalMap.IsOver [X.Over S] [Y.Over S] (f : X ⤏ Y) : Prop where + exists_partialMap_over : ∃ g : X.PartialMap Y, g.IsOver S ∧ g.toRationalMap = f + lemma PartialMap.toRationalMap_surjective : Function.Surjective (@toRationalMap X Y) := Quotient.exists_rep @@ -262,6 +344,14 @@ lemma PartialMap.restrict_toRationalMap (f : X.PartialMap Y) (U : X.Opens) (f.restrict U hU hU').toRationalMap = f.toRationalMap := toRationalMap_eq_iff.mpr (f.restrict_equiv U hU hU') +instance [X.Over S] [Y.Over S] (f : X.PartialMap Y) [f.IsOver S] : f.toRationalMap.IsOver S := + ⟨f, ‹_›, rfl⟩ + +variable (S) in +lemma RationalMap.exists_partialMap_over [X.Over S] [Y.Over S] (f : X ⤏ Y) [f.IsOver S] : + ∃ g : X.PartialMap Y, g.IsOver S ∧ g.toRationalMap = f := + IsOver.exists_partialMap_over + /-- The composition of a rational map and a morphism on the right. -/ def RationalMap.compHom (f : X ⤏ Y) (g : Y ⟶ Z) : X ⤏ Z := by refine Quotient.map (PartialMap.compHom · g) ?_ f @@ -275,6 +365,42 @@ def RationalMap.compHom (f : X ⤏ Y) (g : Y ⟶ Z) : X ⤏ Z := by lemma RationalMap.compHom_toRationalMap (f : X.PartialMap Y) (g : Y ⟶ Z) : (f.compHom g).toRationalMap = f.toRationalMap.compHom g := rfl +instance [X.Over S] [Y.Over S] [Z.Over S] (f : X ⤏ Y) (g : Y ⟶ Z) + [f.IsOver S] [g.IsOver S] : (f.compHom g).IsOver S where + exists_partialMap_over := by + obtain ⟨f, hf, rfl⟩ := f.exists_partialMap_over S + exact ⟨f.compHom g, inferInstance, rfl⟩ + +variable (S) in +lemma PartialMap.exists_restrict_isOver [X.Over S] [Y.Over S] (f : X.PartialMap Y) + [f.toRationalMap.IsOver S] : ∃ U hU hU', (f.restrict U hU hU').IsOver S := by + obtain ⟨f', hf₁, hf₂⟩ := RationalMap.IsOver.exists_partialMap_over (S := S) (f := f.toRationalMap) + obtain ⟨U, hU, hUl, hUr, e⟩ := PartialMap.toRationalMap_eq_iff.mp hf₂ + exact ⟨U, hU, hUr, by rw [IsOver, ← e]; infer_instance⟩ + +lemma RationalMap.isOver_iff [X.Over S] [Y.Over S] {f : X ⤏ Y} : + f.IsOver S ↔ f.compHom (Y ↘ S) = (X ↘ S).toRationalMap := by + constructor + · intro h + obtain ⟨g, hg, e⟩ := f.exists_partialMap_over S + rw [← e, Hom.toRationalMap, ← compHom_toRationalMap, PartialMap.isOver_iff_eq_restrict.mp hg, + PartialMap.restrict_toRationalMap] + · intro e + obtain ⟨f, rfl⟩ := PartialMap.toRationalMap_surjective f + obtain ⟨U, hU, hUl, hUr, e⟩ := PartialMap.toRationalMap_eq_iff.mp e + exact ⟨⟨f.restrict U hU hUl, by simpa using e, by simp⟩⟩ + +lemma PartialMap.isOver_toRationalMap_iff_of_isSeparated [X.Over S] [Y.Over S] [IsReduced X] + [S.IsSeparated] {f : X.PartialMap Y} : + f.toRationalMap.IsOver S ↔ f.IsOver S := by + refine ⟨fun _ ↦ ?_, fun _ ↦ inferInstance⟩ + obtain ⟨U, hU, hU', H⟩ := f.exists_restrict_isOver (S := S) + rw [isOver_iff] + have : IsDominant (X.homOfLE hU') := Opens.isDominant_homOfLE hU _ + exact ext_of_isDominant (ι := X.homOfLE hU') (by simpa using H.1) + +section functionField + /-- A rational map restricts to a map from `Spec K(X)`. -/ noncomputable def RationalMap.fromFunctionField [IrreducibleSpace X] (f : X ⤏ Y) : @@ -328,6 +454,94 @@ def RationalMap.equivFunctionField [IsIntegral X] [LocallyOfFiniteType sY] : (ofFunctionField sX sY f.1.fromFunctionField _) f (RationalMap.fromFunctionField_ofFunctionField _ _ _ _)) +/-- +Given `S`-schemes `X` and `Y` such that `Y` is locally of finite type and `X` is integral, +`S`-morphisms `Spec K(X) ⟶ Y` correspond bijectively to `S`-rational maps from `X` to `Y`. +-/ +noncomputable +def RationalMap.equivFunctionFieldOver [X.Over S] [Y.Over S] [IsIntegral X] + [LocallyOfFiniteType (Y ↘ S)] : + { f : Spec X.functionField ⟶ Y // f.IsOver S } ≃ { f : X ⤏ Y // f.IsOver S } := + ((Equiv.subtypeEquivProp (by simp only [Hom.isOver_iff]; rfl)).trans + (RationalMap.equivFunctionField (X ↘ S) (Y ↘ S))).trans + (Equiv.subtypeEquivProp (by ext f; rw [RationalMap.isOver_iff])) + +end functionField + +section domain + +/-- The domain of definition of a rational map. -/ +def RationalMap.domain (f : X ⤏ Y) : X.Opens := + sSup { PartialMap.domain g | (g) (_ : g.toRationalMap = f) } + +lemma PartialMap.le_domain_toRationalMap (f : X.PartialMap Y) : + f.domain ≤ f.toRationalMap.domain := + le_sSup ⟨f, rfl, rfl⟩ + +lemma RationalMap.mem_domain {f : X ⤏ Y} {x} : + x ∈ f.domain ↔ ∃ g : X.PartialMap Y, x ∈ g.domain ∧ g.toRationalMap = f := + TopologicalSpace.Opens.mem_sSup.trans (by simp [@and_comm (x ∈ _)]) + +lemma RationalMap.dense_domain (f : X ⤏ Y) : Dense (X := X) f.domain := + f.inductionOn (fun g ↦ g.dense_domain.mono g.le_domain_toRationalMap) + +/-- The open cover of the domain of `f : X ⤏ Y`, +consisting of all the domains of the partial maps in the equivalence class. -/ +noncomputable +def RationalMap.openCoverDomain (f : X ⤏ Y) : f.domain.toScheme.OpenCover where + J := { PartialMap.domain g | (g) (_ : g.toRationalMap = f) } + obj U := U.1.toScheme + map U := X.homOfLE (le_sSup U.2) + f x := ⟨_, (TopologicalSpace.Opens.mem_sSup.mp x.2).choose_spec.1⟩ + covers x := ⟨⟨x.1, (TopologicalSpace.Opens.mem_sSup.mp x.2).choose_spec.2⟩, Subtype.ext (by simp)⟩ + +/-- If `f : X ⤏ Y` is a rational map from a reduced scheme to a separated scheme, +then `f` can be represented as a partial map on its domain of definition. -/ +noncomputable +def RationalMap.toPartialMap [IsReduced X] [Y.IsSeparated] (f : X ⤏ Y) : X.PartialMap Y := by + refine ⟨f.domain, f.dense_domain, f.openCoverDomain.glueMorphisms + (fun x ↦ (X.isoOfEq x.2.choose_spec.2).inv ≫ x.2.choose.hom) ?_⟩ + intros x y + let g (x : f.openCoverDomain.J) := x.2.choose + have hg₁ (x) : (g x).toRationalMap = f := x.2.choose_spec.1 + have hg₂ (x) : (g x).domain = x.1 := x.2.choose_spec.2 + refine (cancel_epi (isPullback_opens_inf_le (le_sSup x.2) (le_sSup y.2)).isoPullback.hom).mp ?_ + simp only [openCoverDomain, IsPullback.isoPullback_hom_fst_assoc, + IsPullback.isoPullback_hom_snd_assoc] + show _ ≫ _ ≫ (g x).hom = _ ≫ _ ≫ (g y).hom + simp_rw [← cancel_epi (X.isoOfEq congr($(hg₂ x) ⊓ $(hg₂ y))).hom, ← Category.assoc] + convert (PartialMap.equiv_iff_of_isSeparated (S := ⊤_ _) (f := g x) (g := g y)).mp ?_ using 1 + · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] + · dsimp; congr 1; simp [← cancel_mono (Opens.ι _)] + · rw [← PartialMap.toRationalMap_eq_iff, hg₁, hg₁] + +lemma PartialMap.toPartialMap_toRationalMap_restrict [IsReduced X] [Y.IsSeparated] + (f : X.PartialMap Y) : (f.toRationalMap.toPartialMap.restrict _ f.dense_domain + f.le_domain_toRationalMap).hom = f.hom := by + dsimp [RationalMap.toPartialMap] + refine (f.toRationalMap.openCoverDomain.ι_glueMorphisms _ _ ⟨_, f, rfl, rfl⟩).trans ?_ + generalize_proofs _ _ H _ + have : H.choose = f := (equiv_iff_of_domain_eq_of_isSeparated (S := ⊤_ _) H.choose_spec.2).mp + (toRationalMap_eq_iff.mp H.choose_spec.1) + exact ((ext_iff _ _).mp this.symm).choose_spec.symm + +@[simp] +lemma RationalMap.toRationalMap_toPartialMap [IsReduced X] [Y.IsSeparated] + (f : X ⤏ Y) : f.toPartialMap.toRationalMap = f := by + obtain ⟨f, rfl⟩ := PartialMap.toRationalMap_surjective f + trans (f.toRationalMap.toPartialMap.restrict _ + f.dense_domain f.le_domain_toRationalMap).toRationalMap + · simp + · congr 1 + exact PartialMap.ext _ f rfl (by simpa using f.toPartialMap_toRationalMap_restrict) + +instance [IsReduced X] [Y.IsSeparated] [S.IsSeparated] [X.Over S] [Y.Over S] + (f : X ⤏ Y) [f.IsOver S] : f.toPartialMap.IsOver S := by + rw [← PartialMap.isOver_toRationalMap_iff_of_isSeparated, f.toRationalMap_toPartialMap] + infer_instance + +end domain + end Scheme end AlgebraicGeometry diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index e1e6db017d998..4977b2dfaeffc 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -52,8 +52,10 @@ def ι : ↑U ⟶ X := X.ofRestrict _ instance : IsOpenImmersion U.ι := inferInstanceAs (IsOpenImmersion (X.ofRestrict _)) -@[simps] instance : U.toScheme.Over X := ⟨U.ι⟩ @[simps! over] instance : U.toScheme.CanonicallyOver X where + hom := U.ι + +instance (U : X.Opens) : U.ι.IsOver X where lemma toScheme_carrier : (U : Type u) = (U : Set X) := rfl @@ -202,6 +204,8 @@ lemma Scheme.homOfLE_ι (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : X.homOfLE e ≫ V.ι = U.ι := IsOpenImmersion.lift_fac _ _ _ +instance {U V : X.Opens} (h : U ≤ V) : (X.homOfLE h).IsOver X where + @[simp] lemma Scheme.homOfLE_rfl (X : Scheme.{u}) (U : X.Opens) : X.homOfLE (refl U) = 𝟙 _ := by rw [← cancel_mono U.ι, Scheme.homOfLE_ι, Category.id_comp] @@ -444,6 +448,21 @@ theorem isPullback_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Open -- Porting note: changed `rw` to `erw` erw [pullbackRestrictIsoRestrict_inv_fst]; rw [Category.comp_id] +lemma isPullback_opens_inf_le {X : Scheme} {U V W : X.Opens} (hU : U ≤ W) (hV : V ≤ W) : + IsPullback (X.homOfLE inf_le_left) (X.homOfLE inf_le_right) (X.homOfLE hU) (X.homOfLE hV) := by + refine (isPullback_morphismRestrict (X.homOfLE hV) (W.ι ⁻¹ᵁ U)).of_iso (V.ι.isoImage _ ≪≫ + X.isoOfEq ?_) (W.ι.isoImage _ ≪≫ X.isoOfEq ?_) (Iso.refl _) (Iso.refl _) ?_ ?_ ?_ ?_ + · rw [← TopologicalSpace.Opens.map_comp_obj, ← Scheme.comp_base, Scheme.homOfLE_ι] + exact V.functor_map_eq_inf U + · exact (W.functor_map_eq_inf U).trans (by simpa) + all_goals { simp [← cancel_mono (Scheme.Opens.ι _)] } + +lemma isPullback_opens_inf {X : Scheme} (U V : X.Opens) : + IsPullback (X.homOfLE inf_le_left) (X.homOfLE inf_le_right) U.ι V.ι := + (isPullback_morphismRestrict V.ι U).of_iso (V.ι.isoImage _ ≪≫ X.isoOfEq + (V.functor_map_eq_inf U)) (Iso.refl _) (Iso.refl _) (Iso.refl _) (by simp [← cancel_mono U.ι]) + (by simp [← cancel_mono V.ι]) (by simp) (by simp) + @[simp] lemma morphismRestrict_id {X : Scheme.{u}} (U : X.Opens) : 𝟙 X ∣_ U = 𝟙 _ := by rw [← cancel_mono U.ι, morphismRestrict_ι, Category.comp_id, Category.id_comp] diff --git a/Mathlib/CategoryTheory/Comma/OverClass.lean b/Mathlib/CategoryTheory/Comma/OverClass.lean index bfe80651c0003..c9ad20a1fd78e 100644 --- a/Mathlib/CategoryTheory/Comma/OverClass.lean +++ b/Mathlib/CategoryTheory/Comma/OverClass.lean @@ -31,7 +31,7 @@ universe v u variable {C : Type u} [Category.{v} C] -variable {X Y : C} (f : X ⟶ Y) (S S' : C) +variable {X Y Z : C} (f : X ⟶ Y) (S S' : C) /-- `OverClass X S` is the typeclass containing the data of a structure morphism `X ↘ S : X ⟶ S`. @@ -68,27 +68,39 @@ def CanonicallyOverClass.Simps.over (X S : C) [CanonicallyOverClass X S] : X ⟶ initialize_simps_projections CanonicallyOverClass (hom → over) @[simps] -instance (priority := 100) : OverClass X X := ⟨𝟙 _⟩ +instance : OverClass X X := ⟨𝟙 _⟩ -@[simps] +instance : IsIso (S ↘ S) := inferInstanceAs (IsIso (𝟙 S)) + +-- This cannot be a simp lemma be cause it loops with `comp_over`. +@[simps (config := .lemmasOnly)] instance (priority := 900) [CanonicallyOverClass X Y] [OverClass Y S] : OverClass X S := ⟨X ↘ Y ≫ Y ↘ S⟩ /-- Given `OverClass X S` and `OverClass Y S` and `f : X ⟶ Y`, `HomIsOver f S` is the typeclass asserting `f` commutes with the structure morphisms. -/ class HomIsOver (f : X ⟶ Y) (S : C) [OverClass X S] [OverClass Y S] : Prop where - comp_over : f ≫ Y ↘ S = X ↘ S := by simp + comp_over : f ≫ Y ↘ S = X ↘ S := by aesop @[reassoc (attr := simp)] lemma comp_over [OverClass X S] [OverClass Y S] [HomIsOver f S] : f ≫ Y ↘ S = X ↘ S := HomIsOver.comp_over +instance [OverClass X S] : HomIsOver (𝟙 X) S where + +instance [OverClass X S] [OverClass Y S] [OverClass Z S] + (f : X ⟶ Y) (g : Y ⟶ Z) [HomIsOver f S] [HomIsOver g S] : + HomIsOver (f ≫ g) S where + /-- `Scheme.IsOverTower X Y S` is the typeclass asserting that the structure morphisms `X ↘ Y`, `Y ↘ S`, and `X ↘ S` commute. -/ abbrev IsOverTower (X Y S : C) [OverClass X S] [OverClass Y S] [OverClass X Y] := HomIsOver (X ↘ Y) S +instance [OverClass X S] : IsOverTower X X S where +instance [OverClass X S] : IsOverTower X S S where + instance [CanonicallyOverClass X Y] [OverClass Y S] : IsOverTower X Y S := ⟨rfl⟩ From de26fff24458b9bcab999c26f961ecbc29f36c2d Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Sat, 23 Nov 2024 20:25:40 +0000 Subject: [PATCH 089/250] feat(NumberTheory/LSeries/PrimesInAP): more properties of von Mangoldt on residue classes (#19418) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This continues the sequence of PRs on the way to **Dirichlet's Theorem**. This one adds some properties of the von Mangoldt function `Λ` restricted to a residue class: if further restricted to primes, its support is the set of primes in that residue class, and the function `n ↦ Λ n / n`, restriced to non-primes in a residue class, is summable (the latter will be important in deducing information on the former). See [here](https://leanprover.zulipchat.com/#narrow/channel/144837-PR-reviews/topic/Prerequisites.20for.20PNT.20and.20Dirichlet's.20Thm/near/483904451) on Zulip. --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 102 +++++++++++++++++++ 1 file changed, 102 insertions(+) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 0740c4659750a..6a882afa27f44 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -109,6 +109,108 @@ lemma abscissaOfAbsConv_residueClass_le_one : · simp +contextual only [term, Set.mem_setOf_eq, hn, not_false_eq_true, Set.indicator_of_not_mem, ofReal_zero, zero_div, ite_self] +/-- The set we are interested in (prime numbers in the residue class `a`) is the same as the support +of `ArithmeticFunction.vonMangoldt.residueClass` restricted to primes (and divided by `n`; +this is how this result is used later). -/ +lemma support_residueClass_prime_div : + Function.support (fun n : ℕ ↦ (if n.Prime then residueClass a n else 0) / n) = + {p : ℕ | p.Prime ∧ (p : ZMod q) = a} := by + simp only [Function.support, ne_eq, div_eq_zero_iff, ite_eq_right_iff, + Set.indicator_apply_eq_zero, Set.mem_setOf_eq, Nat.cast_eq_zero, not_or, Classical.not_imp] + ext1 p + simp only [Set.mem_setOf_eq] + exact ⟨fun H ↦ ⟨H.1.1, H.1.2.1⟩, + fun H ↦ ⟨⟨H.1, H.2, vonMangoldt_ne_zero_iff.mpr H.1.isPrimePow⟩, H.1.ne_zero⟩⟩ + +private noncomputable def F₀ (n : ℕ) : ℝ := (if n.Prime then 0 else vonMangoldt n) / n + +private noncomputable def F' (pk : Nat.Primes × ℕ) : ℝ := F₀ (pk.1 ^ (pk.2 + 1)) + +private noncomputable def F'' : Nat.Primes × ℕ → ℝ := F' ∘ (Prod.map _root_.id (· + 1)) + +private lemma F''_le (p : Nat.Primes) (k : ℕ) : F'' (p, k) ≤ 2 * (p : ℝ)⁻¹ ^ (k + 3 / 2 : ℝ) := + calc _ + _ = Real.log p * (p : ℝ)⁻¹ ^ (k + 2) := by + simp only [F'', Function.comp_apply, F', F₀, Prod.map_apply, id_eq, le_add_iff_nonneg_left, + zero_le, Nat.Prime.not_prime_pow, ↓reduceIte, vonMangoldt_apply_prime p.prop, + vonMangoldt_apply_pow (Nat.zero_ne_add_one _).symm, Nat.cast_pow, div_eq_mul_inv, + inv_pow (p : ℝ) (k + 2)] + _ ≤ (p: ℝ) ^ (1 / 2 : ℝ) / (1 / 2) * (p : ℝ)⁻¹ ^ (k + 2) := + mul_le_mul_of_nonneg_right (Real.log_le_rpow_div p.val.cast_nonneg one_half_pos) + (pow_nonneg (inv_nonneg_of_nonneg (Nat.cast_nonneg ↑p)) (k + 2)) + _ = 2 * (p : ℝ)⁻¹ ^ (-1 / 2 : ℝ) * (p : ℝ)⁻¹ ^ (k + 2) := by + simp only [← div_mul, div_one, mul_comm, neg_div, Real.inv_rpow p.val.cast_nonneg, + ← Real.rpow_neg p.val.cast_nonneg, neg_neg] + _ = _ := by + rw [mul_assoc, ← Real.rpow_natCast, + ← Real.rpow_add <| by have := p.prop.pos; positivity, Nat.cast_add, Nat.cast_two, + add_comm, add_assoc] + norm_num + +open Nat.Primes + +private lemma summable_F'' : Summable F'' := by + have hp₀ (p : Nat.Primes) : 0 < (p : ℝ)⁻¹ := inv_pos_of_pos (Nat.cast_pos.mpr p.prop.pos) + have hp₁ (p : Nat.Primes) : (p : ℝ)⁻¹ < 1 := + (inv_lt_one₀ <| mod_cast p.prop.pos).mpr <| Nat.one_lt_cast.mpr <| p.prop.one_lt + suffices Summable fun (pk : Nat.Primes × ℕ) ↦ (pk.1 : ℝ)⁻¹ ^ (pk.2 + 3 / 2 : ℝ) by + refine (Summable.mul_left 2 this).of_nonneg_of_le (fun pk ↦ ?_) (fun pk ↦ F''_le pk.1 pk.2) + simp only [F'', Function.comp_apply, F', F₀, Prod.map_fst, id_eq, Prod.map_snd, Nat.cast_pow] + have := vonMangoldt_nonneg (n := (pk.1 : ℕ) ^ (pk.2 + 2)) + positivity + conv => enter [1, pk]; rw [Real.rpow_add <| hp₀ pk.1, Real.rpow_natCast] + refine (summable_prod_of_nonneg (fun _ ↦ by positivity)).mpr ⟨(fun p ↦ ?_), ?_⟩ + · dsimp only -- otherwise the `exact` below times out + exact Summable.mul_right _ <| summable_geometric_of_lt_one (hp₀ p).le (hp₁ p) + · dsimp only + conv => enter [1, p]; rw [tsum_mul_right, tsum_geometric_of_lt_one (hp₀ p).le (hp₁ p)] + refine (summable_rpow.mpr (by norm_num : -(3 / 2 : ℝ) < -1)).mul_left 2 + |>.of_nonneg_of_le (fun p ↦ ?_) (fun p ↦ ?_) + · have := sub_pos.mpr (hp₁ p) + positivity + · rw [Real.inv_rpow p.val.cast_nonneg, Real.rpow_neg p.val.cast_nonneg] + gcongr + rw [inv_le_comm₀ (sub_pos.mpr (hp₁ p)) zero_lt_two, le_sub_comm, + show (1 : ℝ) - 2⁻¹ = 2⁻¹ by norm_num, inv_le_inv₀ (mod_cast p.prop.pos) zero_lt_two] + exact Nat.ofNat_le_cast.mpr p.prop.two_le + +/-- The function `n ↦ Λ n / n`, restriced to non-primes in a residue class, is summable. +This is used to convert results on `ArithmeticFunction.vonMangoldt.residueClass` to results +on primes in an arithmetic progression. -/ +lemma summable_residueClass_non_primes_div : + Summable fun n : ℕ ↦ (if n.Prime then 0 else residueClass a n) / n := by + have h₀ (n : ℕ) : 0 ≤ (if n.Prime then 0 else residueClass a n) / n := by + have := residueClass_nonneg a n + positivity + have hleF₀ (n : ℕ) : (if n.Prime then 0 else residueClass a n) / n ≤ F₀ n := by + refine div_le_div_of_nonneg_right ?_ n.cast_nonneg + split_ifs; exacts [le_rfl, residueClass_le a n] + refine Summable.of_nonneg_of_le h₀ hleF₀ ?_ + have hF₀ (p : Nat.Primes) : F₀ p.val = 0 := by + simp only [p.prop, ↓reduceIte, zero_div, F₀] + refine (summable_subtype_iff_indicator (s := {n | IsPrimePow n}).mp ?_).congr + fun n ↦ Set.indicator_apply_eq_self.mpr fun (hn : ¬ IsPrimePow n) ↦ ?_ + swap + · simp +contextual only [div_eq_zero_iff, ite_eq_left_iff, vonMangoldt_eq_zero_iff, hn, + not_false_eq_true, implies_true, Nat.cast_eq_zero, true_or, F₀] + have hFF' : + F₀ ∘ Subtype.val (p := fun n ↦ n ∈ {n | IsPrimePow n}) = F' ∘ ⇑prodNatEquiv.symm := by + refine (Equiv.eq_comp_symm prodNatEquiv (F₀ ∘ Subtype.val) F').mpr ?_ + ext1 n + simp only [Function.comp_apply, F'] + congr + rw [hFF'] + refine (Nat.Primes.prodNatEquiv.symm.summable_iff (f := F')).mpr ?_ + have hF'₀ (p : Nat.Primes) : F' (p, 0) = 0 := by simp only [zero_add, pow_one, hF₀, F'] + have hF'₁ : F'' = F' ∘ (Prod.map _root_.id (· + 1)) := by + ext1 + simp only [Function.comp_apply, Prod.map_fst, id_eq, Prod.map_snd, F'', F'] + refine (Function.Injective.summable_iff ?_ fun u hu ↦ ?_).mp <| hF'₁ ▸ summable_F'' + · exact Function.Injective.prodMap (fun ⦃a₁ a₂⦄ a ↦ a) <| add_left_injective 1 + · simp only [Set.range_prod_map, Set.range_id, Set.mem_prod, Set.mem_univ, Set.mem_range, + Nat.exists_add_one_eq, true_and, not_lt, nonpos_iff_eq_zero] at hu + rw [← hF'₀ u.1, ← hu] + variable [NeZero q] {a} /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination From 9b9e6ffdfca324d29cf397ddf8a9cb8151a04240 Mon Sep 17 00:00:00 2001 From: David Loeffler Date: Sat, 23 Nov 2024 20:34:16 +0000 Subject: [PATCH 090/250] feat(NumberTheory/Padics): Mahler coeffs tend to 0 (#19340) This adds the key technical lemma in the proof of Mahler's theorem characterising continuous functions on Zp: for any continuous function, the iterated forward differences at 0 tend to zero. Co-authored-by: Giulio Caflisch --- Mathlib/NumberTheory/Padics/MahlerBasis.lean | 149 ++++++++++++++++++- docs/references.bib | 13 ++ 2 files changed, 159 insertions(+), 3 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/MahlerBasis.lean b/Mathlib/NumberTheory/Padics/MahlerBasis.lean index 543eaf0533d7d..a6be858db5335 100644 --- a/Mathlib/NumberTheory/Padics/MahlerBasis.lean +++ b/Mathlib/NumberTheory/Padics/MahlerBasis.lean @@ -3,6 +3,7 @@ Copyright (c) 2024 David Loeffler. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Giulio Caflisch, David Loeffler -/ +import Mathlib.Algebra.Group.ForwardDiff import Mathlib.Analysis.Normed.Group.Ultra import Mathlib.NumberTheory.Padics.ProperSpace import Mathlib.RingTheory.Binomial @@ -15,15 +16,26 @@ import Mathlib.Topology.ContinuousMap.Compact In this file we introduce the Mahler basis function `mahler k`, for `k : ℕ`, which is the unique continuous map `ℤ_[p] → ℚ_[p]` agreeing with `n ↦ n.choose k` for `n ∈ ℕ`. -In future PR's, we will show that these functions give a Banach basis for the space of continuous -maps `ℤ_[p] → ℚ_[p]`. +We also show that for any continuous function `f` on `ℤ_[p]` (valued in a `p`-adic normed space), +the iterated forward differences `Δ^[n] f 0` tend to 0. For this, we follow the argument of +Bojanić [bojanic74]. + +In future PR's, we will show that the Mahler functions give a Banach basis for the space of +continuous maps `ℤ_[p] → ℚ_[p]`, with the basis coefficients of `f` given by the forward differences +`Δ^[n] f 0`. ## References +* [R. Bojanić, *A simple proof of Mahler's theorem on approximation of continuous functions of a + p-adic variable by polynomials*][bojanic74] * [P. Colmez, *Fonctions d'une variable p-adique*][colmez2010] + +## Tags + +Bojanic -/ -open Finset +open Finset fwdDiff IsUltrametricDist NNReal Filter Topology variable {p : ℕ} [hp : Fact p.Prime] @@ -95,3 +107,134 @@ The uniform norm of the `k`-th Mahler basis function is 1, for every `k`. · -- Show norm 1 is attained at `x = k` refine (le_of_eq ?_).trans ((mahler k).norm_coe_le_norm k) rw [mahler_natCast_eq, Nat.choose_self, Nat.cast_one, norm_one] + +section fwdDiff + +variable {M G : Type*} + +/-- Bound for iterated forward differences of a continuous function from a compact space to a +nonarchimedean seminormed group. -/ +lemma IsUltrametricDist.norm_fwdDiff_iter_apply_le [TopologicalSpace M] [CompactSpace M] + [AddCommMonoid M] [SeminormedAddCommGroup G] [IsUltrametricDist G] + (h : M) (f : C(M, G)) (m : M) (n : ℕ) : ‖Δ_[h]^[n] f m‖ ≤ ‖f‖ := by + -- A proof by induction on `n` would be possible but would involve some messing around to + -- define `Δ_[h]` as an operator on continuous maps (not just on bare functions). So instead we + -- use the formula for `Δ_[h]^[n] f` as a sum. + rw [fwdDiff_iter_eq_sum_shift] + refine norm_sum_le_of_forall_le_of_nonneg (norm_nonneg f) fun i _ ↦ ?_ + exact (norm_zsmul_le _ _).trans (f.norm_coe_le_norm _) + +/-- First step in Bojanić's proof of Mahler's theorem (equation (10) of [bojanic74]): rewrite +`Δ^[n + R] f 0` in a shape that makes it easy to bound `p`-adically. -/ +private lemma bojanic_mahler_step1 [AddCommMonoidWithOne M] [AddCommGroup G] (f : M → G) + (n : ℕ) {R : ℕ} (hR : 1 ≤ R) : + Δ_[1]^[n + R] f 0 = -∑ j ∈ range (R - 1), R.choose (j + 1) • Δ_[1]^[n + (j + 1)] f 0 + + ∑ k ∈ range (n + 1), ((-1 : ℤ) ^ (n - k) * n.choose k) • (f (k + R) - f k) := by + have aux : Δ_[1]^[n + R] f 0 = R.choose (R - 1 + 1) • Δ_[1]^[n + R] f 0 := by + rw [Nat.sub_add_cancel hR, Nat.choose_self, one_smul] + rw [neg_add_eq_sub, eq_sub_iff_add_eq, add_comm, aux, (by omega : n + R = (n + ((R - 1) + 1))), + ← sum_range_succ, Nat.sub_add_cancel hR, + ← sub_eq_iff_eq_add.mpr (sum_range_succ' (fun x ↦ R.choose x • Δ_[1]^[n + x] f 0) R), add_zero, + Nat.choose_zero_right, one_smul] + have : ∑ k ∈ Finset.range (R + 1), R.choose k • Δ_[1]^[n + k] f 0 = Δ_[1]^[n] f R := by + simpa only [← Function.iterate_add_apply, add_comm, nsmul_one, add_zero] using + (shift_eq_sum_fwdDiff_iter 1 (Δ_[1]^[n] f) R 0).symm + simp only [this, fwdDiff_iter_eq_sum_shift (1 : M) f n, mul_comm, nsmul_one, mul_smul, add_comm, + add_zero, smul_sub, sum_sub_distrib] + +end fwdDiff + +namespace PadicInt + +section norm_fwdDiff + +variable {p : ℕ} [hp : Fact p.Prime] {E : Type*} + [NormedAddCommGroup E] [NormedSpace ℚ_[p] E] [IsUltrametricDist E] + +/-- +Second step in Bojanić's proof of Mahler's theorem (equation (11) of [bojanic74]): show that values +`Δ_[1]^[n + p ^ t] f 0` for large enough `n` are bounded by the max of `(‖f‖ / p ^ s)` and `1 / p` +times a sup over values for smaller `n`. + +We use `nnnorm`s on the RHS since `Finset.sup` requires an order with a bottom element. +-/ +private lemma bojanic_mahler_step2 {f : C(ℤ_[p], E)} {s t : ℕ} + (hst : ∀ x y : ℤ_[p], ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s) (n : ℕ) : + ‖Δ_[1]^[n + p ^ t] f 0‖ ≤ max ↑((Finset.range (p ^ t - 1)).sup + fun j ↦ ‖Δ_[1]^[n + (j + 1)] f 0‖₊ / p) (‖f‖ / p ^ s) := by + -- Use previous lemma to rewrite in a convenient form. + rw [bojanic_mahler_step1 _ _ (one_le_pow₀ hp.out.one_le)] + -- Now use ultrametric property and bound each term separately. + refine (norm_add_le_max _ _).trans (max_le_max ?_ ?_) + · -- Bounding the sum over `range (p ^ t - 1)`: every term involves a value `Δ_[1]^[·] f 0` and + -- a binomial coefficient which is divisible by `p` + rw [norm_neg, ← coe_nnnorm, coe_le_coe] + refine nnnorm_sum_le_of_forall_le (fun i hi ↦ Finset.le_sup_of_le hi ?_) + rw [mem_range] at hi + rw [← Nat.cast_smul_eq_nsmul ℚ_[p], nnnorm_smul, div_eq_inv_mul] + refine mul_le_mul_of_nonneg_right ?_ (by simp only [zero_le]) + -- remains to show norm of binomial coeff is `≤ p⁻¹` + have : 0 < (p ^ t).choose (i + 1) := Nat.choose_pos (by omega) + rw [← zpow_neg_one, ← coe_le_coe, coe_nnnorm, Padic.norm_eq_pow_val (mod_cast this.ne'), + coe_zpow, NNReal.coe_natCast, (zpow_right_strictMono₀ (mod_cast hp.out.one_lt)).le_iff_le, + neg_le_neg_iff, Padic.valuation_natCast, Nat.one_le_cast] + exact one_le_padicValNat_of_dvd this <| hp.out.dvd_choose_pow (by omega) (by omega) + · -- Bounding the sum over `range (n + 1)`: every term is small by the choice of `t` + refine norm_sum_le_of_forall_le_of_nonempty nonempty_range_succ (fun i _ ↦ ?_) + calc ‖((-1 : ℤ) ^ (n - i) * n.choose i) • (f (i + ↑(p ^ t)) - f i)‖ + _ ≤ ‖f (i + ↑(p ^ t)) - f i‖ := by + rw [← Int.cast_smul_eq_zsmul ℚ_[p], norm_smul] + apply mul_le_of_le_one_left (norm_nonneg _) + simpa only [← coe_intCast] using norm_le_one _ + _ ≤ ‖f‖ / p ^ s := by + apply hst + rw [Nat.cast_pow, add_sub_cancel_left, norm_pow, norm_p, inv_pow, zpow_neg, zpow_natCast] + +/-- +Explicit bound for the decay rate of the Mahler coefficients of a continuous function on `ℤ_[p]`. +This will be used to prove Mahler's theorem. + -/ +lemma fwdDiff_iter_le_of_forall_le {f : C(ℤ_[p], E)} {s t : ℕ} + (hst : ∀ x y : ℤ_[p], ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s) (n : ℕ) : + ‖Δ_[1]^[n + s * p ^ t] f 0‖ ≤ ‖f‖ / p ^ s := by + -- We show the following more general statement by induction on `k`: + suffices ∀ {k : ℕ}, k ≤ s → ‖Δ_[1]^[n + k * p ^ t] f 0‖ ≤ ‖f‖ / p ^ k from this le_rfl + intro k hk + induction' k with k IH generalizing n + · -- base case just says that `‖Δ^[·] (⇑f) 0‖` is bounded by `‖f‖` + simpa only [zero_mul, pow_zero, add_zero, div_one] using norm_fwdDiff_iter_apply_le 1 f 0 n + · -- induction is the "step 2" lemma above + rw [add_mul, one_mul, ← add_assoc] + refine (bojanic_mahler_step2 hst (n + k * p ^ t)).trans (max_le ?_ ?_) + · rw [← coe_nnnorm, ← NNReal.coe_natCast, ← NNReal.coe_pow, ← NNReal.coe_div, NNReal.coe_le_coe] + refine Finset.sup_le fun j _ ↦ ?_ + rw [pow_succ, ← div_div, div_le_div_iff_of_pos_right (mod_cast hp.out.pos), add_right_comm] + exact_mod_cast IH (n + (j + 1)) (by omega) + · exact div_le_div_of_nonneg_left (norm_nonneg _) + (mod_cast pow_pos hp.out.pos _) (mod_cast pow_le_pow_right₀ hp.out.one_le hk) + +/-- Key lemma for Mahler's theorem: for `f` a continuous function on `ℤ_[p]`, the sequence +`n ↦ Δ^[n] f 0` tends to 0. See `PadicInt.fwdDiff_iter_le_of_forall_le` for an explicit +estimate of the decay rate. -/ +lemma fwdDiff_tendsto_zero (f : C(ℤ_[p], E)) : Tendsto (Δ_[1]^[·] f 0) atTop (𝓝 0) := by + -- first extract an `s` + refine NormedAddCommGroup.tendsto_nhds_zero.mpr (fun ε hε ↦ ?_) + have : Tendsto (fun s ↦ ‖f‖ / p ^ s) _ _ := tendsto_const_nhds.div_atTop + (tendsto_pow_atTop_atTop_of_one_lt (mod_cast hp.out.one_lt)) + obtain ⟨s, hs⟩ := (this.eventually_lt_const hε).exists + refine .mp ?_ (.of_forall fun x hx ↦ lt_of_le_of_lt hx hs) + -- use uniform continuity to find `t` + obtain ⟨t, ht⟩ : ∃ t : ℕ, ∀ x y, ‖x - y‖ ≤ p ^ (-t : ℤ) → ‖f x - f y‖ ≤ ‖f‖ / p ^ s := by + rcases eq_or_ne f 0 with rfl | hf + · -- silly case : f = 0 + simp + have : 0 < ‖f‖ / p ^ s := div_pos (norm_pos_iff.mpr hf) (mod_cast pow_pos hp.out.pos _) + obtain ⟨δ, hδpos, hδf⟩ := f.uniform_continuity _ this + obtain ⟨t, ht⟩ := PadicInt.exists_pow_neg_lt p hδpos + exact ⟨t, fun x y hxy ↦ by simpa only [dist_eq_norm_sub] using (hδf (hxy.trans_lt ht)).le⟩ + filter_upwards [eventually_ge_atTop (s * p ^ t)] with m hm + simpa only [Nat.sub_add_cancel hm] using fwdDiff_iter_le_of_forall_le ht (m - s * p ^ t) + +end norm_fwdDiff + +end PadicInt diff --git a/docs/references.bib b/docs/references.bib index f227b4270fd17..3d890cec41f2a 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -423,6 +423,19 @@ @Book{ bogachev2007 url = {https://doi.org/10.1007/978-3-540-34514-5} } +@Article{ bojanic74, + author = {Bojani\'c, Ranko}, + title = {A simple proof of {M}ahler's theorem on approximation of + continuous functions of a p-adic variable by polynomials}, + journal = {Journal of Number Theory}, + volume = {6}, + pages = {412-415}, + year = {1974}, + issn = {0022-314X}, + doi = {10.1016/0022-314X(74)90038-9}, + url = {https://doi.org/10.1016/0022-314X(74)90038-9} +} + @Book{ bollobas1986, author = {Bollob\'{a}s, B\'{e}la}, title = {Combinatorics: Set Systems, Hypergraphs, Families of From b5155f3e85865ffe02cfd495803e84e98ddab21f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sat, 23 Nov 2024 22:52:24 +0000 Subject: [PATCH 091/250] chore: more simpNF cleanup (#19395) --- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 24 +------------------ Mathlib/NumberTheory/Modular.lean | 3 --- .../ModularForms/SlashActions.lean | 6 ++++- .../NumberTheory/Padics/PadicIntegers.lean | 16 ++++++------- Mathlib/NumberTheory/Padics/RingHoms.lean | 2 +- Mathlib/Order/Closure.lean | 2 +- Mathlib/Order/Compare.lean | 6 ++--- Mathlib/Order/Filter/Germ/Basic.lean | 2 +- Mathlib/Order/Filter/Pointwise.lean | 19 ++++----------- Mathlib/Order/Heyting/Basic.lean | 11 ++++----- Mathlib/Order/Interval/Basic.lean | 1 - Mathlib/Order/LiminfLimsup.lean | 6 ++--- Mathlib/Order/SupIndep.lean | 14 +---------- .../RingTheory/DedekindDomain/SInteger.lean | 1 - Mathlib/RingTheory/Ideal/Operations.lean | 4 +--- Mathlib/RingTheory/Valuation/Basic.lean | 2 +- .../Valuation/ValuationSubring.lean | 8 ++----- 17 files changed, 35 insertions(+), 92 deletions(-) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index bef753458eb59..96de9639e3cfd 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -1719,7 +1719,6 @@ theorem curryFinFinset_symm_apply {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = m <| finSumEquivOfFinset hk hl (Sum.inr i) := rfl --- @[simp] -- Porting note: simpNF linter, lhs simplifies, added aux version below theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) @@ -1734,18 +1733,6 @@ theorem curryFinFinset_symm_apply_piecewise_const {k l n : ℕ} {s : Finset (Fin rw [finSumEquivOfFinset_inr, Finset.piecewise_eq_of_not_mem] exact Finset.mem_compl.1 (Finset.orderEmbOfFin_mem _ _ _) -@[simp] -theorem curryFinFinset_symm_apply_piecewise_const_aux {k l n : ℕ} {s : Finset (Fin n)} - (hk : #s = k) (hl : #sᶜ = l) - (f : MultilinearMap R (fun _ : Fin k => M') (MultilinearMap R (fun _ : Fin l => M') M₂)) - (x y : M') : - ((⇑f fun _ => x) (fun i => (Finset.piecewise s (fun _ => x) (fun _ => y) - ((sᶜ.orderEmbOfFin hl) i))) = f (fun _ => x) fun _ => y) := by - have := curryFinFinset_symm_apply_piecewise_const hk hl f x y - simp only [curryFinFinset_symm_apply, finSumEquivOfFinset_inl, Finset.orderEmbOfFin_mem, - Finset.piecewise_eq_of_mem, finSumEquivOfFinset_inr] at this - exact this - @[simp] theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) @@ -1753,23 +1740,14 @@ theorem curryFinFinset_symm_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : (x : M') : ((curryFinFinset R M₂ M' hk hl).symm f fun _ => x) = f (fun _ => x) fun _ => x := rfl --- @[simp] -- Porting note: simpNF, lhs simplifies, added aux version below theorem curryFinFinset_apply_const {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : (curryFinFinset R M₂ M' hk hl f (fun _ => x) fun _ => y) = f (s.piecewise (fun _ => x) fun _ => y) := by + -- Porting note: `rw` fails refine (curryFinFinset_symm_apply_piecewise_const hk hl _ _ _).symm.trans ?_ - -- `rw` fails rw [LinearEquiv.symm_apply_apply] -@[simp] -theorem curryFinFinset_apply_const_aux {k l n : ℕ} {s : Finset (Fin n)} (hk : #s = k) - (hl : #sᶜ = l) (f : MultilinearMap R (fun _ : Fin n => M') M₂) (x y : M') : - (f fun i => Sum.elim (fun _ => x) (fun _ => y) ((⇑ (Equiv.symm (finSumEquivOfFinset hk hl))) i)) - = f (s.piecewise (fun _ => x) fun _ => y) := by - rw [← curryFinFinset_apply] - apply curryFinFinset_apply_const - end MultilinearMap end Currying diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index ff3eef6b03cda..7eb5dbb5ffec4 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -180,9 +180,6 @@ def lcRow0Extend {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : rw [neg_sq] exact hcd.sq_add_sq_ne_zero, LinearEquiv.refl ℝ (Fin 2 → ℝ)] --- `simpNF` times out, but only in CI where all of `Mathlib` is imported -attribute [nolint simpNF] lcRow0Extend_apply lcRow0Extend_symm_apply - /-- The map `lcRow0` is proper, that is, preimages of cocompact sets are finite in `[[* , *], [c, d]]`. -/ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : diff --git a/Mathlib/NumberTheory/ModularForms/SlashActions.lean b/Mathlib/NumberTheory/ModularForms/SlashActions.lean index bc52891b5dffc..ef9d1a0bcb146 100644 --- a/Mathlib/NumberTheory/ModularForms/SlashActions.lean +++ b/Mathlib/NumberTheory/ModularForms/SlashActions.lean @@ -143,10 +143,14 @@ theorem is_invariant_const (A : SL(2, ℤ)) (x : ℂ) : zpow_neg, zpow_one, inv_one, mul_one, neg_zero, zpow_zero] /-- The constant function 1 is invariant under any element of `SL(2, ℤ)`. -/ --- @[simp] -- Porting note: simpNF says LHS simplifies to something more complex theorem is_invariant_one (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ)] A = (1 : ℍ → ℂ) := is_invariant_const _ _ +/-- Variant of `is_invariant_one` with the left hand side in simp normal form. -/ +@[simp] +theorem is_invariant_one' (A : SL(2, ℤ)) : (1 : ℍ → ℂ) ∣[(0 : ℤ)] (A : GL(2, ℝ)⁺) = 1 := by + simpa using is_invariant_one A + /-- A function `f : ℍ → ℂ` is slash-invariant, of weight `k ∈ ℤ` and level `Γ`, if for every matrix `γ ∈ Γ` we have `f(γ • z)= (c*z+d)^k f(z)` where `γ= ![![a, b], ![c, d]]`, and it acts on `ℍ` via Möbius transformations. -/ diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 53e32586436a1..f9e0e0463f74b 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -153,8 +153,7 @@ def Coe.ringHom : ℤ_[p] →+* ℚ_[p] := (subring p).subtype @[simp, norm_cast] theorem coe_pow (x : ℤ_[p]) (n : ℕ) : (↑(x ^ n) : ℚ_[p]) = (↑x : ℚ_[p]) ^ n := rfl --- @[simp] -- Porting note: not in simpNF -theorem mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := Subtype.coe_eta _ _ +theorem mk_coe (k : ℤ_[p]) : (⟨k, k.2⟩ : ℤ_[p]) = k := by simp /-- The inverse of a `p`-adic integer with norm equal to `1` is also a `p`-adic integer. Otherwise, the inverse is defined to be `0`. -/ @@ -165,10 +164,8 @@ instance : CharZero ℤ_[p] where cast_injective m n h := Nat.cast_injective (R := ℚ_[p]) (by rw [Subtype.ext_iff] at h; norm_cast at h) -@[norm_cast] -- @[simp] -- Porting note: not in simpNF -theorem intCast_eq (z1 z2 : ℤ) : (z1 : ℤ_[p]) = z2 ↔ z1 = z2 := by - suffices (z1 : ℚ_[p]) = z2 ↔ z1 = z2 from Iff.trans (by norm_cast) this - norm_cast +@[norm_cast] +theorem intCast_eq (z1 z2 : ℤ) : (z1 : ℤ_[p]) = z2 ↔ z1 = z2 := by simp @[deprecated (since := "2024-04-05")] alias coe_int_eq := intCast_eq @@ -280,8 +277,7 @@ theorem norm_eq_padic_norm {q : ℚ_[p]} (hq : ‖q‖ ≤ 1) : @norm ℤ_[p] _ @[simp] theorem norm_p : ‖(p : ℤ_[p])‖ = (p : ℝ)⁻¹ := padicNormE.norm_p --- @[simp] -- Porting note: not in simpNF -theorem norm_p_pow (n : ℕ) : ‖(p : ℤ_[p]) ^ n‖ = (p : ℝ) ^ (-n : ℤ) := padicNormE.norm_p_pow n +theorem norm_p_pow (n : ℕ) : ‖(p : ℤ_[p]) ^ n‖ = (p : ℝ) ^ (-n : ℤ) := by simp private def cauSeq_to_rat_cauSeq (f : CauSeq ℤ_[p] norm) : CauSeq ℚ_[p] fun a => ‖a‖ := ⟨fun n => f n, fun _ hε => by simpa [norm, norm_def] using f.cauchy hε⟩ @@ -413,10 +409,12 @@ theorem norm_lt_one_mul {z1 z2 : ℤ_[p]} (hz2 : ‖z2‖ < 1) : ‖z1 * z2‖ < ‖z1 * z2‖ = ‖z1‖ * ‖z2‖ := by simp _ < 1 := mul_lt_one_of_nonneg_of_lt_one_right (norm_le_one _) (norm_nonneg _) hz2 --- @[simp] -- Porting note: not in simpNF theorem mem_nonunits {z : ℤ_[p]} : z ∈ nonunits ℤ_[p] ↔ ‖z‖ < 1 := by rw [lt_iff_le_and_ne]; simp [norm_le_one z, nonunits, isUnit_iff] +theorem not_isUnit_iff {z : ℤ_[p]} : ¬IsUnit z ↔ ‖z‖ < 1 := by + simpa using mem_nonunits + /-- A `p`-adic number `u` with `‖u‖ = 1` is a unit of `ℤ_[p]`. -/ def mkUnits {u : ℚ_[p]} (h : ‖u‖ = 1) : ℤ_[p]ˣ := let z : ℤ_[p] := ⟨u, le_of_eq h⟩ diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index dab80d6e19bd6..b80610a74c52a 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -393,7 +393,7 @@ theorem ker_toZModPow (n : ℕ) : rw [zmod_congr_of_sub_mem_span n x _ 0 _ h, cast_zero] apply appr_spec --- @[simp] -- Porting note: not in simpNF +-- This is not a simp lemma; simp can't match the LHS. theorem zmod_cast_comp_toZModPow (m n : ℕ) (h : m ≤ n) : (ZMod.castHom (pow_dvd_pow p h) (ZMod (p ^ m))).comp (@toZModPow p _ n) = @toZModPow p _ m := by apply ZMod.ringHom_eq_of_ker_eq diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index b9c15b5db30cc..8fb5b69f0275c 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -387,7 +387,7 @@ variable [PartialOrder α] [PartialOrder β] {u : β → α} (l : LowerAdjoint u theorem mem_closed_iff_closure_le (x : α) : x ∈ l.closed ↔ u (l x) ≤ x := l.closureOperator.isClosed_iff_closure_le -@[simp, nolint simpNF] -- Porting note: lemma does prove itself, seems to be a linter error +@[simp] theorem closure_is_closed (x : α) : u (l x) ∈ l.closed := l.idempotent x diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index 9fd8b1100315f..5e29b6e29a069 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -137,8 +137,7 @@ theorem cmp_swap [Preorder α] [@DecidableRel α (· < ·)] (a b : α) : (cmp a by_cases h : a < b <;> by_cases h₂ : b < a <;> simp [h, h₂, Ordering.swap] exact lt_asymm h h₂ --- Porting note: Not sure why the simpNF linter doesn't like this. @semorrison -@[simp, nolint simpNF] +@[simp] theorem cmpLE_toDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : α) : cmpLE (toDual x) (toDual y) = cmpLE y x := rfl @@ -148,8 +147,7 @@ theorem cmpLE_ofDual [LE α] [@DecidableRel α (· ≤ ·)] (x y : αᵒᵈ) : cmpLE (ofDual x) (ofDual y) = cmpLE y x := rfl --- Porting note: Not sure why the simpNF linter doesn't like this. @semorrison -@[simp, nolint simpNF] +@[simp] theorem cmp_toDual [LT α] [@DecidableRel α (· < ·)] (x y : α) : cmp (toDual x) (toDual y) = cmp y x := rfl diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 21eda99605828..00156fd9037ce 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -237,7 +237,7 @@ theorem coe_compTendsto (f : α → β) {lc : Filter γ} {g : γ → α} (hg : T rfl -- Porting note https://github.com/leanprover-community/mathlib4/issues/10959 --- simp cannot prove this +-- simp can't match the LHS. @[simp, nolint simpNF] theorem compTendsto'_coe (f : Germ l β) {lc : Filter γ} {g : γ → α} (hg : Tendsto g lc l) : f.compTendsto' _ hg.germ_tendsto = f.compTendsto g hg := diff --git a/Mathlib/Order/Filter/Pointwise.lean b/Mathlib/Order/Filter/Pointwise.lean index c76d14970e05a..05f1a4c1cf4e0 100644 --- a/Mathlib/Order/Filter/Pointwise.lean +++ b/Mathlib/Order/Filter/Pointwise.lean @@ -146,9 +146,8 @@ theorem pureOneHom_apply (a : α) : pureOneHom a = pure a := variable [One β] @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. protected theorem map_one [FunLike F α β] [OneHomClass F α β] (φ : F) : map φ 1 = 1 := by - rw [Filter.map_one', map_one, pure_one] + simp end One @@ -303,9 +302,7 @@ theorem mul_pure : f * pure b = f.map (· * b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_mul_pure : (pure a : Filter α) * pure b = pure (a * b) := - map₂_pure +theorem pure_mul_pure : (pure a : Filter α) * pure b = pure (a * b) := by simp @[to_additive (attr := simp)] theorem le_mul_iff : h ≤ f * g ↔ ∀ ⦃s⦄, s ∈ f → ∀ ⦃t⦄, t ∈ g → s * t ∈ h := @@ -408,9 +405,7 @@ theorem div_pure : f / pure b = f.map (· / b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_div_pure : (pure a : Filter α) / pure b = pure (a / b) := - map₂_pure +theorem pure_div_pure : (pure a : Filter α) / pure b = pure (a / b) := by simp @[to_additive] protected theorem div_le_div : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ / g₁ ≤ f₂ / g₂ := @@ -826,9 +821,7 @@ theorem smul_pure : f • pure b = f.map (· • b) := map₂_pure_right @[to_additive] --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_smul_pure : (pure a : Filter α) • (pure b : Filter β) = pure (a • b) := - map₂_pure +theorem pure_smul_pure : (pure a : Filter α) • (pure b : Filter β) = pure (a • b) := by simp @[to_additive] theorem smul_le_smul : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ • g₁ ≤ f₂ • g₂ := @@ -914,9 +907,7 @@ theorem pure_vsub : (pure a : Filter β) -ᵥ g = g.map (a -ᵥ ·) := theorem vsub_pure : f -ᵥ pure b = f.map (· -ᵥ b) := map₂_pure_right --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11119): removed `simp` attribute because `simpNF` says it can prove it. -theorem pure_vsub_pure : (pure a : Filter β) -ᵥ pure b = (pure (a -ᵥ b) : Filter α) := - map₂_pure +theorem pure_vsub_pure : (pure a : Filter β) -ᵥ pure b = (pure (a -ᵥ b) : Filter α) := by simp theorem vsub_le_vsub : f₁ ≤ f₂ → g₁ ≤ g₂ → f₁ -ᵥ g₁ ≤ f₂ -ᵥ g₂ := map₂_mono diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index 4c1f777c64b8f..c44de8d3c3d14 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -1084,11 +1084,11 @@ theorem top_eq : (⊤ : PUnit) = unit := theorem bot_eq : (⊥ : PUnit) = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem sup_eq : a ⊔ b = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem inf_eq : a ⊓ b = unit := rfl @@ -1096,16 +1096,15 @@ theorem inf_eq : a ⊓ b = unit := theorem compl_eq : aᶜ = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem sdiff_eq : a \ b = unit := rfl -@[simp, nolint simpNF] +@[simp] theorem hnot_eq : ¬a = unit := rfl --- eligible for `dsimp` -@[simp, nolint simpNF] +@[simp] theorem himp_eq : a ⇨ b = unit := rfl diff --git a/Mathlib/Order/Interval/Basic.lean b/Mathlib/Order/Interval/Basic.lean index 014a33ce871f3..55171625e65d7 100644 --- a/Mathlib/Order/Interval/Basic.lean +++ b/Mathlib/Order/Interval/Basic.lean @@ -109,7 +109,6 @@ theorem mem_mk {hx : x.1 ≤ x.2} : a ∈ mk x hx ↔ x.1 ≤ a ∧ a ≤ x.2 := theorem mem_def : a ∈ s ↔ s.fst ≤ a ∧ a ≤ s.snd := Iff.rfl --- @[simp] -- Porting note: not in simpNF theorem coe_nonempty (s : NonemptyInterval α) : (s : Set α).Nonempty := nonempty_Icc.2 s.fst_le_snd diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 25404163646dc..2576f34d60217 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -829,15 +829,13 @@ theorem HasBasis.limsup_eq_sInf_univ_of_empty {f : ι → α} {v : Filter ι} limsup f v = sInf univ := HasBasis.liminf_eq_sSup_univ_of_empty (α := αᵒᵈ) hv i hi h'i --- Porting note: simp_nf linter incorrectly says: lhs does not simplify when using simp on itself. -@[simp, nolint simpNF] +@[simp] theorem liminf_nat_add (f : ℕ → α) (k : ℕ) : liminf (fun i => f (i + k)) atTop = liminf f atTop := by change liminf (f ∘ (· + k)) atTop = liminf f atTop rw [liminf, liminf, ← map_map, map_add_atTop_eq_nat] --- Porting note: simp_nf linter incorrectly says: lhs does not simplify when using simp on itself. -@[simp, nolint simpNF] +@[simp] theorem limsup_nat_add (f : ℕ → α) (k : ℕ) : limsup (fun i => f (i + k)) atTop = limsup f atTop := @liminf_nat_add αᵒᵈ _ f k diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 7e9d4b4cdbb85..68e516ec2adec 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -180,19 +180,7 @@ theorem SupIndep.attach (hs : s.SupIndep f) : s.attach.SupIndep fun a => f a := obtain ⟨j, hj, hji⟩ := hi' rwa [Subtype.ext hji] at hj -/- -Porting note: simpNF linter returns - -"Left-hand side does not simplify, when using the simp lemma on itself." - -However, simp does indeed solve the following. - -example {α ι} [Lattice α] [OrderBot α] (s : Finset ι) (f : ι → α) : - (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by simp - -See https://github.com/leanprover-community/batteries/issues/71 for the simpNF issue. --/ -@[simp, nolint simpNF] +@[simp] theorem supIndep_attach : (s.attach.SupIndep fun a => f a) ↔ s.SupIndep f := by refine ⟨fun h t ht i his hit => ?_, SupIndep.attach⟩ classical diff --git a/Mathlib/RingTheory/DedekindDomain/SInteger.lean b/Mathlib/RingTheory/DedekindDomain/SInteger.lean index 83b4fe6cf9021..7568c704d46d4 100644 --- a/Mathlib/RingTheory/DedekindDomain/SInteger.lean +++ b/Mathlib/RingTheory/DedekindDomain/SInteger.lean @@ -96,7 +96,6 @@ theorem unit_valuation_eq_one (x : S.unit K) {v : HeightOneSpectrum R} (hv : v v.valuation ((x : Kˣ) : K) = 1 := x.property v hv --- Porting note: `apply_inv_coe` fails the simpNF linter /-- The group of `S`-units is the group of units of the ring of `S`-integers. -/ @[simps apply_val_coe symm_apply_coe] def unitEquivUnitsInteger : S.unit K ≃* (S.integer K)ˣ where diff --git a/Mathlib/RingTheory/Ideal/Operations.lean b/Mathlib/RingTheory/Ideal/Operations.lean index 13dfba254995d..7218ec0a9688e 100644 --- a/Mathlib/RingTheory/Ideal/Operations.lean +++ b/Mathlib/RingTheory/Ideal/Operations.lean @@ -287,9 +287,7 @@ variable {R : Type u} [Semiring R] theorem add_eq_sup {I J : Ideal R} : I + J = I ⊔ J := rfl --- dsimp loops when applying this lemma to its LHS, --- probably https://github.com/leanprover/lean4/pull/2867 -@[simp, nolint simpNF] +@[simp] theorem zero_eq_bot : (0 : Ideal R) = ⊥ := rfl diff --git a/Mathlib/RingTheory/Valuation/Basic.lean b/Mathlib/RingTheory/Valuation/Basic.lean index e4f462f138500..fe79b5797d6c9 100644 --- a/Mathlib/RingTheory/Valuation/Basic.lean +++ b/Mathlib/RingTheory/Valuation/Basic.lean @@ -128,7 +128,7 @@ theorem coe_mk (f : R →*₀ Γ₀) (h) : ⇑(Valuation.mk f h) = f := rfl theorem toFun_eq_coe (v : Valuation R Γ₀) : v.toFun = v := rfl -@[simp] -- Porting note: requested by simpNF as toFun_eq_coe LHS simplifies +@[simp] theorem toMonoidWithZeroHom_coe_eq_coe (v : Valuation R Γ₀) : (v.toMonoidWithZeroHom : R → Γ₀) = v := rfl diff --git a/Mathlib/RingTheory/Valuation/ValuationSubring.lean b/Mathlib/RingTheory/Valuation/ValuationSubring.lean index 867c268d02ac5..813f07bdbde31 100644 --- a/Mathlib/RingTheory/Valuation/ValuationSubring.lean +++ b/Mathlib/RingTheory/Valuation/ValuationSubring.lean @@ -50,9 +50,6 @@ instance : SetLike (ValuationSubring K) K where replace h := SetLike.coe_injective' h congr --- Porting note https://github.com/leanprover-community/mathlib4/issues/10959 --- simp cannot prove this -@[simp, nolint simpNF] theorem mem_carrier (x : K) : x ∈ A.carrier ↔ x ∈ A := Iff.refl _ @[simp] @@ -668,9 +665,9 @@ def unitsModPrincipalUnitsEquivResidueFieldUnits : local instance : MulOneClass ({ x // x ∈ unitGroup A } ⧸ Subgroup.comap (Subgroup.subtype (unitGroup A)) (principalUnitGroup A)) := inferInstance --- @[simp] -- Porting note: not in simpNF theorem unitsModPrincipalUnitsEquivResidueFieldUnits_comp_quotientGroup_mk : - A.unitsModPrincipalUnitsEquivResidueFieldUnits.toMonoidHom.comp (QuotientGroup.mk' _) = + (A.unitsModPrincipalUnitsEquivResidueFieldUnits : _ ⧸ Subgroup.comap _ _ →* _).comp + (QuotientGroup.mk' (A.principalUnitGroup.subgroupOf A.unitGroup)) = A.unitGroupToResidueFieldUnits := rfl theorem unitsModPrincipalUnitsEquivResidueFieldUnits_comp_quotientGroup_mk_apply @@ -793,7 +790,6 @@ namespace Valuation variable {Γ : Type*} [LinearOrderedCommGroupWithZero Γ] (v : Valuation K Γ) (x : Kˣ) --- @[simp] -- Porting note: not in simpNF theorem mem_unitGroup_iff : x ∈ v.valuationSubring.unitGroup ↔ v x = 1 := IsEquiv.eq_one_iff_eq_one (Valuation.isEquiv_valuation_valuationSubring _).symm From 1051dd99dda0730db5d0f1ede40137f1c75e9742 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 24 Nov 2024 10:25:30 +1100 Subject: [PATCH 092/250] . --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Group/Action/Defs.lean | 8 +- Mathlib/Algebra/Group/AddChar.lean | 2 +- Mathlib/Algebra/Group/Equiv/Basic.lean | 20 +++- Mathlib/Algebra/Module/LinearMap/Defs.lean | 4 +- Mathlib/Algebra/Order/Field/Power.lean | 4 +- .../Algebra/Order/Group/Pointwise/Bounds.lean | 7 +- Mathlib/Analysis/Complex/Circle.lean | 6 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 8 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/Limits/Creates.lean | 12 +-- .../Limits/FunctorCategory/Basic.lean | 8 +- Mathlib/Data/Complex/Abs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/LazyList/Basic.lean | 10 +- Mathlib/Data/List/Basic.lean | 8 +- Mathlib/Data/List/Flatten.lean | 3 +- Mathlib/Data/List/Indexes.lean | 12 +-- Mathlib/Data/List/Lex.lean | 2 +- Mathlib/Data/List/Pairwise.lean | 3 +- Mathlib/Data/List/Permutation.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 9 +- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Deprecated/AlgebraClasses.lean | 47 +++++----- Mathlib/Deprecated/ByteArray.lean | 10 +- Mathlib/Deprecated/Combinator.lean | 6 +- Mathlib/Deprecated/Equiv.lean | 4 +- Mathlib/Deprecated/Logic.lean | 74 +++++++-------- Mathlib/Deprecated/NatLemmas.lean | 8 +- Mathlib/Deprecated/RelClasses.lean | 8 +- Mathlib/GroupTheory/Coset/Basic.lean | 8 +- Mathlib/GroupTheory/Coset/Defs.lean | 21 ++++- Mathlib/GroupTheory/Exponent.lean | 2 +- Mathlib/GroupTheory/GroupAction/Defs.lean | 8 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 17 +++- Mathlib/GroupTheory/OrderOfElement.lean | 8 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 9 +- Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 10 +- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 20 ++-- Mathlib/LinearAlgebra/Prod.lean | 28 +++--- Mathlib/Logic/Basic.lean | 9 +- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Measure/WithDensityFinite.lean | 2 +- Mathlib/NumberTheory/ArithmeticFunction.lean | 2 +- Mathlib/NumberTheory/MulChar/Basic.lean | 6 +- Mathlib/Order/Defs.lean | 8 +- Mathlib/Order/OmegaCompletePartialOrder.lean | 2 +- Mathlib/Order/RelClasses.lean | 6 +- Mathlib/SetTheory/Cardinal/Aleph.lean | 18 ++-- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 94 +++++++++---------- Mathlib/SetTheory/Ordinal/Basic.lean | 8 +- Mathlib/SetTheory/Ordinal/Enum.lean | 2 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 66 ++++++------- Mathlib/SetTheory/Ordinal/Principal.lean | 4 +- Mathlib/SetTheory/ZFC/Basic.lean | 42 ++++----- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- Mathlib/Topology/Algebra/Group/Quotient.lean | 6 +- Mathlib/Topology/LocalAtTarget.lean | 4 +- Mathlib/Topology/MetricSpace/Polish.lean | 2 +- Mathlib/Topology/NoetherianSpace.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 2 +- lake-manifest.json | 2 +- 66 files changed, 406 insertions(+), 321 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index 7bc0dd9d0dc2e..7dc9b9e5ba7f8 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -634,7 +634,7 @@ end MonoidHom end MonoidHom set_option linter.deprecated false in -@[simp, deprecated "Deprecated without replacement." (since := "2024-10-17")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-10-17")] lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl namespace List diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 823dbf69f3586..294281be04be5 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -297,9 +297,15 @@ lemma smul_mul_smul_comm [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] (a • b) * (c • d) = (a * c) • (b * d) := by have : SMulCommClass β α β := .symm ..; exact smul_smul_smul_comm a b c d -@[to_additive (attr := deprecated (since := "2024-08-29"))] +@[to_additive] alias smul_mul_smul := smul_mul_smul_comm +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated smul_mul_smul_comm (since := "2024-08-29")] smul_mul_smul +attribute [deprecated vadd_add_vadd_comm (since := "2024-08-29")] vadd_add_vadd + + /-- Note that the `IsScalarTower α β β` and `SMulCommClass α β β` typeclass arguments are usually satisfied by `Algebra α β`. -/ @[to_additive] diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index e66bd4f51f989..0a1e982e42b01 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -253,7 +253,7 @@ lemma ne_one_iff : ψ ≠ 1 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff lemma ne_zero_iff : ψ ≠ 0 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff /-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-06")] +@[deprecated "No deprecation message was provided." (since := "2024-06-06")] def IsNontrivial (ψ : AddChar A M) : Prop := ∃ a : A, ψ a ≠ 1 set_option linter.deprecated false in diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 81aff3106d93e..ea5ec7e66d31d 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -125,12 +125,28 @@ class MulEquivClass (F : Type*) (A B : outParam Type*) [Mul A] [Mul B] [EquivLik /-- Preserves multiplication. -/ map_mul : ∀ (f : F) (a b), f (a * b) = f a * f b -@[to_additive (attr := deprecated (since := "2024-11-10"))] +@[to_additive] alias MulEquivClass.map_eq_one_iff := EmbeddingLike.map_eq_one_iff -@[to_additive (attr := deprecated (since := "2024-11-10"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_eq_one_iff (since := "2024-11-10")] +MulEquivClass.map_eq_one_iff +attribute [deprecated EmbeddingLike.map_eq_zero_iff (since := "2024-11-10")] +AddEquivClass.map_eq_zero_iff + + +@[to_additive] alias MulEquivClass.map_ne_one_iff := EmbeddingLike.map_ne_one_iff +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_ne_one_iff (since := "2024-11-10")] +MulEquivClass.map_ne_one_iff +attribute [deprecated EmbeddingLike.map_ne_zero_iff (since := "2024-11-10")] +AddEquivClass.map_ne_zero_iff + + namespace MulEquivClass variable (F) diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index c6436036fb7c1..d4c25f58c52fb 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -596,14 +596,14 @@ variable [Semiring R] [Module R M] [Semiring S] [Module S M₂] [Module R M₃] variable {σ : R →+* S} /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toSemilinearMap (fₗ : M →ₑ+[σ.toMonoidHom] M₂) : M →ₛₗ[σ] M₂ := { fₗ with } instance : SemilinearMapClass (M →ₑ+[σ.toMonoidHom] M₂) σ M M₂ where /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toLinearMap (fₗ : M →+[R] M₃) : M →ₗ[R] M₃ := { fₗ with } diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index 7df61ba43d34a..9cec9ad4fcc29 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -73,13 +73,13 @@ theorem zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : Injective (a ^ · : ℤ theorem zpow_inj (h₀ : 0 < a) (h₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := zpow_right_inj₀ h₀ h₁ -@[deprecated "Deprecated without replacement." (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_of_min_le {x : α} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) := have : Antitone fun n : ℤ => x ^ (-n) := fun _ _ h => zpow_le_zpow_right₀ hx (neg_le_neg h) (this h).trans_eq this.map_min -@[deprecated "Deprecated without replacement." (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_iff_min_le {x : α} (hx : 1 < x) {a b c : ℤ} : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) ↔ min a b ≤ c := by simp_rw [le_max_iff, min_le_iff, zpow_le_zpow_iff_right₀ hx, neg_le_neg_iff] diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index ceb329876b8e6..294f6be7c92df 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -48,7 +48,12 @@ lemma BddAbove.mul (hs : BddAbove s) (ht : BddAbove t) : BddAbove (s * t) := lemma BddBelow.mul (hs : BddBelow s) (ht : BddBelow t) : BddBelow (s * t) := (Nonempty.mul hs ht).mono (subset_lowerBounds_mul s t) -@[to_additive (attr := deprecated (since := "2024-11-13"))] alias Set.BddAbove.mul := BddAbove.mul +@[to_additive] alias Set.BddAbove.mul := BddAbove.mul + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated BddAbove.mul (since := "2024-11-13")] Set.BddAbove.mul +attribute [deprecated BddAbove.add (since := "2024-11-13")] Set.BddAbove.add @[to_additive] lemma BddAbove.range_mul (hf : BddAbove (range f)) (hg : BddAbove (range g)) : diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index b56e79503ce1d..e4a9a2e283278 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -41,17 +41,17 @@ open ComplexConjugate /-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. Please use `Circle` when referring to the circle as a type. -/ -@[deprecated "Deprecated without replacement." (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] def circle : Submonoid ℂ := Submonoid.unitSphere ℂ set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := mem_sphere_zero_iff_norm set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by simp [Complex.abs, mem_circle_iff_abs] diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index 4f96b555b2c3c..2db2c795ef2c2 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -92,7 +92,7 @@ lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } include adj in -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := adj.leftAdjoint_preservesColimits @@ -117,7 +117,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E := Functor.reflectsColimits_of_isEquivalence E @@ -216,7 +216,7 @@ lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } include adj in -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := adj.rightAdjoint_preservesLimits @@ -240,7 +240,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated "Deprecated without replacement." (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E := Functor.reflectsLimits_of_isEquivalence E diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index a159d404af292..d4b4fc8a4f69e 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -111,7 +111,7 @@ theorem comp_f {P Q R : Karoubi C} (f : P ⟶ Q) (g : Q ⟶ R) : (f ≫ g).f = f @[simp] theorem id_f {P : Karoubi C} : Hom.f (𝟙 P) = P.p := rfl -@[deprecated "Deprecated without replacement." (since := "2024-07-15")] +@[deprecated "No deprecation message was provided." (since := "2024-07-15")] theorem id_eq {P : Karoubi C} : 𝟙 P = ⟨P.p, by repeat' rw [P.idem]⟩ := rfl /-- It is possible to coerce an object of `C` into an object of `Karoubi C`. diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 19ff0630c5b87..9769ea2520000 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -332,7 +332,7 @@ instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 _ _ @@ -342,7 +342,7 @@ lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) instance (priority := 100) preservesLimitOfShape_of_createsLimitsOfShape_and_hasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F where -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F := @@ -354,7 +354,7 @@ instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F := @@ -466,7 +466,7 @@ instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K (Cocones.functoriality K F).mapIso ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 _ _ @@ -477,7 +477,7 @@ instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F := @@ -489,7 +489,7 @@ instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F := diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index a39f5182361aa..603001e9e6ce6 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -362,7 +362,7 @@ lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) exact isLimitOfPreserves _ hc⟩⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -374,7 +374,7 @@ lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Cate PreservesLimitsOfShape J F := ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -387,7 +387,7 @@ lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) ⟨fun {L} _ => preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -412,7 +412,7 @@ lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) exact isColimitOfPreserves _ hc⟩⟩ -@[deprecated "Deprecated without replacement." (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index acbd1c017c03a..18f2f48e1b525 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -199,7 +199,7 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated "Deprecated without replacement." (since := "2024-02-14")] +@[deprecated "No deprecation message was provided." (since := "2024-02-14")] lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index de29683071348..4fcd646907769 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1115,7 +1115,7 @@ lemma succAbove_ne_last {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a ≠ last _) lemma succAbove_last_apply (i : Fin n) : succAbove (last n) i = castSucc i := by rw [succAbove_last] -@[deprecated "Deprecated without replacement." (since := "2024-05-30")] +@[deprecated "No deprecation message was provided." (since := "2024-05-30")] lemma succAbove_lt_ge (p : Fin (n + 1)) (i : Fin n) : castSucc i < p ∨ p ≤ castSucc i := Nat.lt_or_ge (castSucc i) p diff --git a/Mathlib/Data/LazyList/Basic.lean b/Mathlib/Data/LazyList/Basic.lean index 00b68b137a00c..5ba81d38999a8 100644 --- a/Mathlib/Data/LazyList/Basic.lean +++ b/Mathlib/Data/LazyList/Basic.lean @@ -24,7 +24,7 @@ namespace LazyList open Function /-- Isomorphism between strict and lazy lists. -/ -@[deprecated "Deprecated without replacement." (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] def listEquivLazyList (α : Type*) : List α ≃ LazyList α where toFun := LazyList.ofList invFun := LazyList.toList @@ -39,12 +39,12 @@ def listEquivLazyList (α : Type*) : List α ≃ LazyList α where · simp [toList, ofList] · simpa [ofList, toList] -@[deprecated "Deprecated without replacement." (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : Traversable LazyList where map := @LazyList.traverse Id _ traverse := @LazyList.traverse -@[deprecated "Deprecated without replacement." (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulTraversable LazyList := by apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs · induction xs using LazyList.rec with @@ -71,7 +71,7 @@ instance : LawfulTraversable LazyList := by Function.comp_def, Thunk.pure, ofList] | mk _ ih => apply ih -@[deprecated "Deprecated without replacement." (since := "2024-07-22"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-07-22"), simp] theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with | nil => simp [LazyList.bind] @@ -81,7 +81,7 @@ theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by simp [ih] | mk f ih => simp_all -@[deprecated "Deprecated without replacement." (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulMonad LazyList := LawfulMonad.mk' (id_map := by intro α xs diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 8e27205a4ec40..08a1dc0307429 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -34,7 +34,7 @@ universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} /-- `≤` implies not `>` for lists. -/ -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl -- Porting note: Delete this attribute @@ -496,7 +496,7 @@ theorem get_tail (l : List α) (i) (h : i < l.tail.length) l.tail.get ⟨i, h⟩ = l.get ⟨i + 1, h'⟩ := by cases l <;> [cases h; rfl] -@[deprecated (since := "2024-08-22")] +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] theorem get_cons {l : List α} {a : α} {n} (hl) : (a :: l).get ⟨n, hl⟩ = if hn : n = 0 then a else l.get ⟨n - 1, by contrapose! hl; rw [length_cons]; omega⟩ := @@ -627,7 +627,7 @@ lemma cons_sublist_cons' {a b : α} : a :: l₁ <+ b :: l₂ ↔ a :: l₁ <+ l theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ := h.cons _ -@[deprecated (since := "2024-04-07")] +@[deprecated "No deprecation message was provided." (since := "2024-04-07")] theorem sublist_of_cons_sublist_cons {a} (h : a :: l₁ <+ a :: l₂) : l₁ <+ l₂ := h.of_cons_cons @[deprecated (since := "2024-04-07")] alias cons_sublist_cons_iff := cons_sublist_cons @@ -2257,7 +2257,7 @@ theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs simp; omega set_option linter.deprecated false in -@[deprecated (since := "2024-07-25")] +@[deprecated "No deprecation message was provided." (since := "2024-07-25")] theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by induction xs generalizing i j hj with diff --git a/Mathlib/Data/List/Flatten.lean b/Mathlib/Data/List/Flatten.lean index f6fbc8e3044fc..797ba9dc5c297 100644 --- a/Mathlib/Data/List/Flatten.lean +++ b/Mathlib/Data/List/Flatten.lean @@ -158,7 +158,8 @@ theorem append_flatten_map_append (L : List (List α)) (x : List α) : @[deprecated (since := "2024-10-15")] alias append_join_map_append := append_flatten_map_append -@[deprecated (since := "2024-08-15")] theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] +theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : l <+ L.flatten := sublist_flatten_of_mem h diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index 477ac3d84bfa8..36b5dee0789f5 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -84,7 +84,7 @@ theorem mapIdx_eq_ofFn (l : List α) (f : ℕ → α → β) : section deprecated /-- Lean3 `map_with_index` helper function -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β | _, [] => [] | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as @@ -92,12 +92,12 @@ protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List set_option linter.deprecated false in /-- Given a function `f : ℕ → α → β` and `as : List α`, `as = [a₀, a₁, ...]`, returns the list `[f 0 a₀, f 1 a₁, ...]`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := List.oldMapIdxCore f 0 as set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : ℕ) : l.oldMapIdxCore f n = l.oldMapIdx fun i a ↦ f (i + n) a := by induction' l with hd tl hl generalizing f n @@ -106,7 +106,7 @@ protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : simp only [List.oldMapIdxCore, hl, Nat.add_left_comm, Nat.add_comm, Nat.add_zero] set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) (l₁ l₂ : List α), List.oldMapIdxCore f n (l₁ ++ l₂) = List.oldMapIdxCore f n l₁ ++ List.oldMapIdxCore f (n + l₁.length) l₂ := by @@ -134,7 +134,7 @@ protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) ( rw [Nat.add_assoc]; simp only [Nat.add_comm] set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) (e : α), List.oldMapIdx f (l ++ [e]) = List.oldMapIdx f l ++ [f l.length e] := by intros f l e @@ -143,7 +143,7 @@ protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) ( simp only [Nat.zero_add]; rfl set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem new_def_eq_old_def : ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by intro f diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index fef097d96a69f..69566d9aa5d91 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -99,7 +99,7 @@ instance isAsymm (r : α → α → Prop) [IsAsymm α r] : IsAsymm (List α) (Le | _, _, Lex.cons _, Lex.rel h₂ => asymm h₂ h₂ | _, _, Lex.cons h₁, Lex.cons h₂ => aux _ _ h₁ h₂ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictTotalOrder (r : α → α → Prop) [IsStrictTotalOrder α r] : IsStrictTotalOrder (List α) (Lex r) := { isStrictWeakOrder_of_isOrderConnected with } diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 4682916a610a0..be4f87bb84be6 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -50,7 +50,8 @@ theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x hl.forall hr -- Porting note: Duplicate of `pairwise_map` but with `f` explicit. -@[deprecated (since := "2024-02-25")] theorem pairwise_map' (f : β → α) : +@[deprecated "No deprecation message was provided." (since := "2024-02-25")] +theorem pairwise_map' (f : β → α) : ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l | [] => by simp only [map, Pairwise.nil] | b :: l => by diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 849d50190c764..f436730690267 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -419,7 +419,7 @@ theorem length_permutations'Aux (s : List α) (x : α) : · simp · simpa using IH -@[deprecated "Deprecated without replacement." (since := "2024-06-12")] +@[deprecated "No deprecation message was provided." (since := "2024-06-12")] theorem permutations'Aux_get_zero (s : List α) (x : α) (hn : 0 < length (permutations'Aux x s) := (by simp)) : (permutations'Aux x s).get ⟨0, hn⟩ = x :: s := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 9063d971434ff..40fa79abc7b3a 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1085,7 +1085,7 @@ lemma sub_mod_eq_zero_of_mod_eq (h : m % k = n % k) : (m - n) % k = 0 := by lemma one_mod_eq_one : ∀ {n : ℕ}, 1 % n = 1 ↔ n ≠ 1 | 0 | 1 | n + 2 => by simp -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] lemma one_mod_of_ne_one : ∀ {n : ℕ}, n ≠ 1 → 1 % n = 1 := one_mod_eq_one.mpr lemma dvd_sub_mod (k : ℕ) : n ∣ k - k % n := diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 1f4a5da3be5cd..e7008d4c67fbf 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -415,8 +415,13 @@ lemma disjoint_smul_set_left : Disjoint (a • s) t ↔ Disjoint s (a⁻¹ • t lemma disjoint_smul_set_right : Disjoint s (a • t) ↔ Disjoint (a⁻¹ • s) t := by simpa using disjoint_smul_set (a := a) (s := a⁻¹ • s) -@[to_additive (attr := deprecated (since := "2024-10-18"))] -alias smul_set_disjoint_iff := disjoint_smul_set +@[to_additive] alias smul_set_disjoint_iff := disjoint_smul_set + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated disjoint_smul_set (since := "2024-10-18")] smul_set_disjoint_iff +attribute [deprecated disjoint_vadd_set (since := "2024-10-18")] vadd_set_disjoint_iff + end Group diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index d8394f09dc005..bbaa5b239fbfe 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -36,12 +36,12 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-29")] +@[deprecated "No deprecation message was provided." (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index d93789e6ec5dc..bb6f948871913 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -25,16 +25,16 @@ universe u v variable {α : Sort u} -@[deprecated "Deprecated without replacement." (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where left_cancel : ∀ a b c, op a b = op a c → b = c -@[deprecated "Deprecated without replacement." (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsRightCancel (α : Sort u) (op : α → α → α) : Prop where right_cancel : ∀ a b c, op a b = op c b → a = c /-- `IsTotalPreorder X r` means that the binary relation `r` on `X` is total and a preorder. -/ -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsTotalPreorder (α : Sort u) (r : α → α → Prop) extends IsTrans α r, IsTotal α r : Prop /-- Every total pre-order is a pre-order. -/ @@ -45,11 +45,11 @@ instance (priority := 100) isTotalPreorder_isPreorder (α : Sort u) (r : α → /-- `IsIncompTrans X lt` means that for `lt` a binary relation on `X`, the incomparable relation `fun a b => ¬ lt a b ∧ ¬ lt b a` is transitive. -/ -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsIncompTrans (α : Sort u) (lt : α → α → Prop) : Prop where incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeakOrder α lt] : IsIncompTrans α lt := { ‹IsStrictWeakOrder α lt› with } @@ -59,7 +59,7 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := IsIncompTrans.incomp_trans _ _ _ @@ -68,7 +68,8 @@ section ExplicitRelationVariants variable (r) -@[elab_without_expected_type, deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[elab_without_expected_type, + deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans_of [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := incomp_trans @@ -85,32 +86,32 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] def Equiv (a b : α) : Prop := ¬a ≺ b ∧ ¬b ≺ a local infixl:50 " ≈ " => @Equiv _ r -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem esymm {a b : α} : a ≈ b → b ≈ a := fun ⟨h₁, h₂⟩ => ⟨h₂, h₁⟩ -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv {a b : α} : a ≈ b → ¬a ≺ b := fun h => h.1 -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv' {a b : α} : a ≈ b → ¬b ≺ a := fun h => h.2 variable [IsStrictWeakOrder α r] -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem erefl (a : α) : a ≈ a := ⟨irrefl a, irrefl a⟩ -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem etrans {a b c : α} : a ≈ b → b ≈ c → a ≈ c := incomp_trans -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isEquiv : IsEquiv α (@Equiv _ r) where refl := erefl trans _ _ _ := etrans @@ -123,7 +124,7 @@ notation:50 a " ≈[" lt "]" b:50 => @Equiv _ lt a b end StrictWeakOrder -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem isStrictWeakOrder_of_isTotalPreorder {α : Sort u} {le : α → α → Prop} {lt : α → α → Prop} [DecidableRel le] [IsTotalPreorder α le] (h : ∀ a b, lt a b ↔ ¬le b a) : IsStrictWeakOrder α lt := @@ -149,20 +150,20 @@ section LinearOrder variable {α : Type*} [LinearOrder α] set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance : IsTotalPreorder α (· ≤ ·) where trans := @le_trans _ _ total := le_total set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictWeakOrder_of_linearOrder : IsStrictWeakOrder α (· < ·) := have : IsTotalPreorder α (· ≤ ·) := by infer_instance -- Porting note: added isStrictWeakOrder_of_isTotalPreorder lt_iff_not_ge end LinearOrder -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, lt a b → ¬lt b c ∧ ¬lt c b → lt a c := @fun a b c hab ⟨nbc, ncb⟩ => @@ -171,7 +172,7 @@ theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt a b ∧ ¬lt b a := incomp_trans_of lt ⟨nac, nca⟩ ⟨ncb, nbc⟩ absurd hab this.1 -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, ¬lt a b ∧ ¬lt b a → lt b c → lt a c := @fun a b c ⟨nab, nba⟩ hbc => @@ -180,7 +181,7 @@ theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt b c ∧ ¬lt c b := incomp_trans_of lt ⟨nba, nab⟩ ⟨nac, nca⟩ absurd hbc this.1 -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : ¬lt a b ∧ ¬lt b a → a = b := fun ⟨nab, nba⟩ => match trichotomous_of lt a b with @@ -188,17 +189,17 @@ theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α | Or.inr (Or.inl hab) => hab | Or.inr (Or.inr hba) => absurd hba nba -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_eqv_lt {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : a ≈[lt]b → a = b := eq_of_incomp -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : ¬lt a b ∧ ¬lt b a ↔ a = b := Iff.intro eq_of_incomp fun hab => hab ▸ And.intro (irrefl_of lt a) (irrefl_of lt a) -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eqv_lt_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : a ≈[lt]b ↔ a = b := incomp_iff_eq a b diff --git a/Mathlib/Deprecated/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean index 35ac422df4fe2..b7146dbf289c6 100644 --- a/Mathlib/Deprecated/ByteArray.lean +++ b/Mathlib/Deprecated/ByteArray.lean @@ -19,7 +19,7 @@ set_option linter.deprecated false namespace Nat /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def Up (ub a i : Nat) := i < a ∧ i < ub theorem Up.next {ub i} (h : i < ub) : Up ub (i+1) i := ⟨Nat.lt_succ_self _, h⟩ @@ -28,13 +28,13 @@ theorem Up.WF (ub) : WellFounded (Up ub) := Subrelation.wf (h₂ := (measure (ub - ·)).wf) fun ⟨ia, iu⟩ ↦ Nat.sub_lt_sub_left iu ia /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def upRel (ub : Nat) : WellFoundedRelation Nat := ⟨Up ub, Up.WF ub⟩ end Nat /-- A terminal byte slice, a suffix of a byte array. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] structure ByteSliceT := (arr : ByteArray) (off : Nat) namespace ByteSliceT @@ -66,7 +66,7 @@ def toArray : ByteSlice → ByteArray universe u v /-- The inner loop of the `forIn` implementation for byte slices. -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β → m (ForInStep β)) (arr : ByteArray) (off _end : Nat) (i : Nat) (b : β) : m β := if h : i < _end then do @@ -75,7 +75,7 @@ def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β | ForInStep.yield b => have := Nat.Up.next h; loop f arr off _end (i+1) b else pure b -@[deprecated "Deprecated without replacement." (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] instance {m : Type u → Type v} : ForIn m ByteSlice UInt8 := ⟨fun ⟨arr, off, len⟩ b f ↦ forIn.loop f arr off (off + len) off b⟩ diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean index 3c16256244179..366fd5997f982 100644 --- a/Mathlib/Deprecated/Combinator.lean +++ b/Mathlib/Deprecated/Combinator.lean @@ -15,11 +15,11 @@ namespace Combinator universe u v w variable {α : Sort u} {β : Sort v} {γ : Sort w} -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def I (a : α) := a -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def K (a : α) (_b : β) := a -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) end Combinator diff --git a/Mathlib/Deprecated/Equiv.lean b/Mathlib/Deprecated/Equiv.lean index 2491e31e194a1..b31fe578386cc 100644 --- a/Mathlib/Deprecated/Equiv.lean +++ b/Mathlib/Deprecated/Equiv.lean @@ -21,10 +21,10 @@ variable {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α set_option linter.deprecated false -@[deprecated "Deprecated without replacement." (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ -@[deprecated "Deprecated without replacement." (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ diff --git a/Mathlib/Deprecated/Logic.lean b/Mathlib/Deprecated/Logic.lean index ce91b89b2cdf4..52d03620dc96a 100644 --- a/Mathlib/Deprecated/Logic.lean +++ b/Mathlib/Deprecated/Logic.lean @@ -37,43 +37,43 @@ local infix:65 (priority := high) " + " => g def Commutative := ∀ a b, a * b = b * a @[deprecated Std.Associative (since := "2024-09-13")] def Associative := ∀ a b c, (a * b) * c = a * (b * c) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftIdentity := ∀ a, one * a = a -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightIdentity := ∀ a, a * one = a -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightInverse := ∀ a, a * a⁻¹ = one -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftCancelative := ∀ a b c, a * b = a * c → b = c -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightCancelative := ∀ a b c, a * b = c * b → a = c -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftDistributive := ∀ a b c, a * (b + c) = a * b + a * c -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightDistributive := ∀ a b c, (a + b) * c = a * c + b * c end Binary @[deprecated (since := "2024-09-03")] alias not_of_eq_false := of_eq_false -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem cast_proof_irrel {β : Sort u} (h₁ h₂ : α = β) (a : α) : cast h₁ a = cast h₂ a := rfl @[deprecated (since := "2024-09-03")] alias eq_rec_heq := eqRec_heq @[deprecated (since := "2024-09-03")] alias heq_prop := proof_irrel_heq -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_left {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a = a') → (h₂ : Eq.rec (motive := fun a _ ↦ φ a) p₁ e = p₂) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_right {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a' = a) → (h₂ : p₁ = Eq.rec (motive := fun a _ ↦ φ a) p₂ e) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem of_heq_true {a : Prop} (h : HEq a True) : a := of_eq_true (eq_of_heq h) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem eq_rec_compose {α β φ : Sort u} : ∀ (p₁ : β = φ) (p₂ : α = β) (a : α), (Eq.recOn p₁ (Eq.recOn p₂ a : β) : φ) = Eq.recOn (Eq.trans p₂ p₁) a @@ -114,20 +114,20 @@ theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) /- decidable -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_True' (h : Decidable True) : decide True = true := by simp -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_False' (h : Decidable False) : decide False = false := by simp namespace Decidable -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_true [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : p) (h₄ : h₁ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isTrue _ => rfl) h₄ -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : ¬p) (h₄ : h₂ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isFalse _ => rfl) h₄ @@ -145,25 +145,25 @@ end Decidable @[deprecated (since := "2024-09-03")] alias decidableTrue := instDecidableTrue @[deprecated (since := "2024-09-03")] alias decidableFalse := instDecidableFalse -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecEq {α : Sort u} (p : α → α → Bool) : Prop := ∀ ⦃x y : α⦄, p x y = true → x = y -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecRefl {α : Sort u} (p : α → α → Bool) : Prop := ∀ x, p x x = true -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def decidableEq_of_bool_pred {α : Sort u} {p : α → α → Bool} (h₁ : IsDecEq p) (h₂ : IsDecRefl p) : DecidableEq α | x, y => if hp : p x y = true then isTrue (h₁ hp) else isFalse (fun hxy : x = y ↦ absurd (h₂ y) (by rwa [hxy] at hp)) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inl_refl {α : Sort u} [h : DecidableEq α] (a : α) : h a a = isTrue (Eq.refl a) := match h a a with | isTrue _ => rfl -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} (n : a ≠ b) : h a b = isFalse n := match h a b with @@ -171,7 +171,7 @@ theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} /- subsingleton -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} [h₃ : ∀ h : p, Subsingleton (h₁ h)] [h₄ : ∀ h : ¬p, Subsingleton (h₂ h)] : Subsingleton (Decidable.recOn h h₂ h₁) := @@ -179,15 +179,15 @@ theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h | isTrue h => h₃ h | isFalse h => h₄ h -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_pos {c t e : Prop} [Decidable c] (h : ite c t e) (hc : c) : t := (if_pos hc ▸ h :) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_neg {c t e : Prop} [Decidable c] (h : ite c t e) (hnc : ¬c) : e := (if_neg hnc ▸ h :) -@[deprecated "Deprecated without replacement." (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : Decidable c] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -199,7 +199,7 @@ theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : | isFalse h₁, isTrue h₂ => absurd h₂ (Iff.mp (not_congr h_c) h₁) | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : Decidable c] (h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) : ite b x y ↔ ite c u v := match dec_b, dec_c with @@ -209,12 +209,12 @@ theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : De | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -- @[congr] -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_congr_prop {b c x y u v : Prop} [Decidable b] [Decidable c] (h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) : ite b x y ↔ ite c u v := if_ctx_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : c → (x ↔ u)) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -222,7 +222,7 @@ theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c (h_e : ¬c → (y ↔ v)) : ite b x y ↔ @ite _ c (decidable_of_decidable_of_iff h_c) u v := if_ctx_congr_prop (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : x ↔ u) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -230,7 +230,7 @@ theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h (h_e : y ↔ v) : ite b x y ↔ (@ite _ c (decidable_of_decidable_of_iff h_c) u v) := if_ctx_simp_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -241,31 +241,31 @@ theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] dite b x y = @dite _ c (decidable_of_decidable_of_iff h_c) u v := dif_ctx_congr (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] def AsTrue (c : Prop) [Decidable c] : Prop := if c then True else False -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def AsFalse (c : Prop) [Decidable c] : Prop := if c then False else True -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem AsTrue.get {c : Prop} [h₁ : Decidable c] (_ : AsTrue c) : c := match h₁ with | isTrue h_c => h_c /- Equalities for rewriting let-expressions -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_eq {α : Sort u} {β : Sort v} {a₁ a₂ : α} (b : α → β) (h : a₁ = a₂) : (let x : α := a₁; b x) = (let x : α := a₂; b x) := congrArg b h -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_heq {α : Sort v} {β : α → Sort u} {a₁ a₂ : α} (b : ∀ x : α, β x) (h : a₁ = a₂) : HEq (let x : α := a₁; b x) (let x : α := a₂; b x) := by cases h; rfl -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_body_eq {α : Sort v} {β : α → Sort u} (a : α) {b₁ b₂ : ∀ x : α, β x} (h : ∀ x, b₁ x = b₂ x) : (let x : α := a; b₁ x) = (let x : α := a; b₂ x) := by exact h _ ▸ rfl -@[deprecated "Deprecated without replacement." (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_eq {α : Sort v} {β : Sort u} {a₁ a₂ : α} {b₁ b₂ : α → β} (h₁ : a₁ = a₂) (h₂ : ∀ x, b₁ x = b₂ x) : (let x : α := a₁; b₁ x) = (let x : α := a₂; b₂ x) := by simp [h₁, h₂] diff --git a/Mathlib/Deprecated/NatLemmas.lean b/Mathlib/Deprecated/NatLemmas.lean index fa59bd023418a..81b124b4476cd 100644 --- a/Mathlib/Deprecated/NatLemmas.lean +++ b/Mathlib/Deprecated/NatLemmas.lean @@ -28,7 +28,7 @@ namespace Nat /-! successor and predecessor -/ -@[deprecated "Deprecated without replacement." (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ m → B) : B := by induction n with | zero => exact H1 rfl @@ -36,7 +36,7 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem one_eq_succ_zero : 1 = succ 0 := rfl @@ -44,7 +44,7 @@ theorem one_eq_succ_zero : 1 = succ 0 := -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P (succ n) 0) (H3 : ∀ n m, P n m → P (succ n) (succ m)) : ∀ n m : ℕ, P n m | 0, _m => H1 _ @@ -55,7 +55,7 @@ def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cond_decide_mod_two (x : ℕ) [d : Decidable (x % 2 = 1)] : cond (@decide (x % 2 = 1) d) 1 0 = x % 2 := by simp only [cond_eq_if, decide_eq_true_eq] diff --git a/Mathlib/Deprecated/RelClasses.lean b/Mathlib/Deprecated/RelClasses.lean index b9f1af89148f5..8d5f647030d5d 100644 --- a/Mathlib/Deprecated/RelClasses.lean +++ b/Mathlib/Deprecated/RelClasses.lean @@ -27,14 +27,14 @@ variable {α : Type u} open Function -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } -@[deprecated "Deprecated without replacement." (since := "2024-08-22")] +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where -@[deprecated "Deprecated without replacement." (since := "2024-08-22")] +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where -@[deprecated "Deprecated without replacement." (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsIncompTrans α (· < ·) := by infer_instance diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index a93c6f92a2109..4f8c575b67fcf 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -318,9 +318,15 @@ lemma orbit_eq_out_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x 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"))] +@[to_additive] alias orbit_eq_out'_smul := orbit_eq_out_smul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbit_eq_out_smul (since := "2024-10-19")] orbit_eq_out'_smul +attribute [deprecated QuotientAddGroup.orbit_eq_out_vadd (since := "2024-10-19")] +QuotientAddGroup.orbit_eq_out'_vadd + end QuotientGroup namespace Subgroup diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index 50c360a3c7413..27eb12564db66 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -170,8 +170,13 @@ theorem induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (Quoti instance : Coe α (α ⧸ s) := ⟨mk⟩ -@[to_additive (attr := deprecated (since := "2024-08-04"))] -alias induction_on' := induction_on +@[to_additive] alias induction_on' := induction_on + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated induction_on (since := "2024-08-04")] induction_on' +attribute [deprecated QuotientAddGroup.induction_on (since := "2024-08-04")] +QuotientAddGroup.induction_on' @[to_additive (attr := simp)] theorem quotient_liftOn_mk {β} (f : α → β) (h) (x : α) : Quotient.liftOn' (x : α ⧸ s) f h = f x := @@ -195,7 +200,7 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := _ ↔ leftRel s a b := Quotient.eq'' _ ↔ _ := by rw [leftRel_apply] -@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-04"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-04"))] alias eq' := QuotientGroup.eq @[to_additive] @@ -211,10 +216,16 @@ variable (s) 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 "Deprecated without replacement." (since := "2024-10-19")) - QuotientAddGroup.mk_out'_eq_mul] +@[to_additive QuotientAddGroup.mk_out'_eq_mul] alias mk_out'_eq_mul := mk_out_eq_mul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated mk_out_eq_mul (since := "2024-10-19")] mk_out'_eq_mul +attribute [deprecated QuotientAddGroup.mk_out_eq_mul (since := "2024-10-19")] +QuotientAddGroup.mk_out'_eq_mul + + variable {s} {a b : α} @[to_additive (attr := simp)] diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 1d8a38f4b7f43..ea64a5ce67e5e 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -645,7 +645,7 @@ lemma inv_eq_self_of_orderOf_eq_two {x : G} (hx : orderOf x = 2) : -- TODO: delete /-- Any group of exponent two is abelian. -/ @[to_additive (attr := reducible, - deprecated "Deprecated without replacement." (since := "2024-02-17")) + deprecated "No deprecation message was provided." (since := "2024-02-17")) "Any additive group of exponent two is abelian."] def instCommGroupOfExponentTwo (hG : Monoid.exponent G = 2) : CommGroup G where mul_comm := mul_comm_of_exponent_two hG diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index 3848050964715..fa15f5e7d6253 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -312,9 +312,15 @@ variable {G α} theorem orbitRel_apply {a b : α} : orbitRel G α a b ↔ a ∈ orbit G b := Iff.rfl -@[to_additive (attr := deprecated (since := "2024-10-18"))] +@[to_additive] alias orbitRel_r_apply := orbitRel_apply +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbitRel_apply (since := "2024-10-18")] orbitRel_r_apply +attribute [deprecated AddAction.orbitRel_apply (since := "2024-10-18")] AddAction.orbitRel_r_apply + + /-- When you take a set `U` in `α`, push it down to the quotient, and pull back, you get the union of the orbit of `U` under `G`. -/ @[to_additive diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 4c15d5da76105..778fb1159400a 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -103,13 +103,26 @@ theorem _root_.QuotientGroup.out_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ 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"))] +@[to_additive] alias Quotient.mk_smul_out' := Quotient.mk_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.mk_smul_out (since := "2024-10-19")] Quotient.mk_smul_out' +attribute [deprecated AddAction.Quotient.mk_vadd_out (since := "2024-10-19")] +AddAction.Quotient.mk_vadd_out' + -- Porting note: removed simp attribute, simp can prove this -@[to_additive (attr := deprecated (since := "2024-10-19"))] +@[to_additive] alias Quotient.coe_smul_out' := Quotient.coe_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.coe_smul_out (since := "2024-10-19")] Quotient.coe_smul_out' +attribute [deprecated AddAction.Quotient.coe_vadd_out (since := "2024-10-19")] +AddAction.Quotient.coe_vadd_out' + + @[deprecated (since := "2024-10-19")] alias _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem := QuotientGroup.out_conj_pow_minimalPeriod_mem diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index b380c1f71c410..bde4f07b48a62 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -88,9 +88,15 @@ theorem not_isOfFinOrder_of_injective_pow {x : G} (h : Injective fun n : ℕ => theorem IsOfFinOrder.one : IsOfFinOrder (1 : G) := isOfFinOrder_iff_pow_eq_one.mpr ⟨1, Nat.one_pos, one_pow 1⟩ -@[to_additive (attr := deprecated (since := "2024-10-11"))] +@[to_additive] alias isOfFinOrder_one := IsOfFinOrder.one +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated IsOfFinOrder.one (since := "2024-10-11")] isOfFinOrder_one +attribute [deprecated IsOfFinAddOrder.zero (since := "2024-10-11")] isOfFinAddOrder_zero + + @[to_additive] lemma IsOfFinOrder.pow {n : ℕ} : IsOfFinOrder a → IsOfFinOrder (a ^ n) := by simp_rw [isOfFinOrder_iff_pow_eq_one] diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index ca2a56cc761c7..7900709a53790 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -327,10 +327,17 @@ lemma isCyclic_iff_exists_ofOrder_eq_natCard [Finite α] : refine isCyclic_of_orderOf_eq_card g ?_ simp [hg] -@[to_additive (attr := deprecated (since := "2024-04-20"))] +@[to_additive] protected alias IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype := isCyclic_iff_exists_ofOrder_eq_natCard +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated isCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype +attribute [deprecated isAddCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsAddCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype + section variable [Fintype α] diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index 006312d2cecef..da171b8d6bfaa 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -58,7 +58,7 @@ namespace LinearMap namespace BilinForm -@[deprecated "Deprecated without replacement." (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y' | _, _, _, _, rfl, rfl => rfl @@ -101,7 +101,7 @@ theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _ -@[deprecated "Deprecated without replacement." (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := rfl @@ -111,7 +111,7 @@ theorem zero_apply (x y : M) : (0 : BilinForm R M) x y = 0 := variable (B D B₁ D₁) -@[deprecated "Deprecated without replacement." (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_add : ⇑(B + D) = B + D := rfl @@ -119,7 +119,7 @@ theorem coe_add : ⇑(B + D) = B + D := theorem add_apply (x y : M) : (B + D) x y = B x y + D x y := rfl -@[deprecated "Deprecated without replacement." (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_neg : ⇑(-B₁) = -B₁ := rfl @@ -127,7 +127,7 @@ theorem coe_neg : ⇑(-B₁) = -B₁ := theorem neg_apply (x y : M₁) : (-B₁) x y = -B₁ x y := rfl -@[deprecated "Deprecated without replacement." (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_sub : ⇑(B₁ - D₁) = B₁ - D₁ := rfl diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 8fe60684491f7..c1b171f04d13e 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -58,16 +58,16 @@ section ToLin' def toLinHomAux₁ (A : BilinForm R M) (x : M) : M →ₗ[R] R := A x /-- Auxiliary definition to define `toLinHom`; see below. -/ -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHomAux₂ (A : BilinForm R M) : M →ₗ[R] M →ₗ[R] R := A /-- The linear map obtained from a `BilinForm` by fixing the left co-ordinate and evaluating in the right. -/ -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHom : BilinForm R M →ₗ[R] M →ₗ[R] M →ₗ[R] R := LinearMap.id set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem toLin'_apply (A : BilinForm R M) (x : M) : toLinHom (M := M) A x = A x := rfl @@ -112,7 +112,7 @@ def LinearMap.toBilinAux (f : M →ₗ[R] M →ₗ[R] R) : BilinForm R M := f set_option linter.deprecated false in /-- Bilinear forms are linearly equivalent to maps with two arguments that are linear in both. -/ -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] R := { BilinForm.toLinHom with invFun := LinearMap.toBilinAux @@ -121,35 +121,35 @@ def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] set_option linter.deprecated false in /-- A map with two arguments that is linear in both is linearly equivalent to bilinear form. -/ -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.toBilin : (M →ₗ[R] M →ₗ[R] R) ≃ₗ[R] BilinForm R M := BilinForm.toLin.symm -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilinAux_eq (f : M →ₗ[R] M →ₗ[R] R) : LinearMap.toBilinAux f = f := rfl set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_symm : (LinearMap.toBilin.symm : BilinForm R M ≃ₗ[R] _) = BilinForm.toLin := rfl set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_symm : (BilinForm.toLin.symm : _ ≃ₗ[R] BilinForm R M) = LinearMap.toBilin := LinearMap.toBilin.symm_symm set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_apply (f : M →ₗ[R] M →ₗ[R] R) (x y : M) : toBilin f x y = f x y := rfl set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_apply (x : M) : BilinForm.toLin B x = B x := rfl diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 3d4187a8405b4..7bbb4ac2f4b65 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -834,17 +834,17 @@ set_option linter.deprecated false /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i /-- Auxiliary definition for `tunnel`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M | 0 => ⟨⊤, LinearEquiv.ofTop ⊤ rfl⟩ | n + 1 => @@ -855,7 +855,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M)ᵒᵈ := -- Note: the hint `(α := _)` had to be added in https://github.com/leanprover-community/mathlib4/pull/8386 ⟨fun n => OrderDual.toDual (α := Submodule R M) (tunnel' f i n).1, @@ -867,24 +867,24 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) /-- Each `tailing f i n` is a copy of `N`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by dsimp [tailing, tunnelAux] rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by rw [disjoint_iff] @@ -893,7 +893,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by @@ -902,19 +902,19 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv apply Submodule.map_subtype_le /-- The supremum of all the copies of `N` found inside the tunnel. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) -@[simp, deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] -@[simp, deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by induction' n with n ih @@ -926,7 +926,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel -@[deprecated "Deprecated without replacement." (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := Disjoint.mono_right (tailing_le_tunnel f i _) (tailings_disjoint_tunnel f i _) diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 7a0eb1ff4d84f..89c3a8fe8a2f1 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -725,20 +725,21 @@ alias by_contradiction := byContradiction -- TODO: remove? rename in core? alias prop_complete := propComplete -- TODO: remove? rename in core? -@[elab_as_elim, deprecated (since := "2024-07-27")] theorem cases_true_false (p : Prop → Prop) +@[elab_as_elim, deprecated "No deprecation message was provided." (since := "2024-07-27")] +theorem cases_true_false (p : Prop → Prop) (h1 : p True) (h2 : p False) (a : Prop) : p a := Or.elim (prop_complete a) (fun ht : a = True ↦ ht.symm ▸ h1) fun hf : a = False ↦ hf.symm ▸ h2 -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem eq_false_or_eq_true (a : Prop) : a = False ∨ a = True := (prop_complete a).symm set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p a := @cases_true_false p h1 h2 a set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 end Classical diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index 10d9db322ebca..599c6658734e6 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -567,7 +567,7 @@ instance isFiniteMeasure_sfiniteSeq [h : SFinite μ] (n : ℕ) : IsFiniteMeasure h.1.choose_spec.1 n set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] instance isFiniteMeasure_sFiniteSeq [SFinite μ] (n : ℕ) : IsFiniteMeasure (sFiniteSeq μ n) := isFiniteMeasure_sfiniteSeq n diff --git a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean index 2f78a1dd221a8..40b56846328c7 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean @@ -117,7 +117,7 @@ noncomputable def Measure.densityToFinite (μ : Measure α) [SFinite μ] (a : α μ.rnDeriv μ.toFinite a set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-04")] +@[deprecated "No deprecation message was provided." (since := "2024-10-04")] lemma densityToFinite_def (μ : Measure α) [SFinite μ] : μ.densityToFinite = μ.rnDeriv μ.toFinite := rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 77a36164d0f5b..8b8038b188ab0 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -1291,7 +1291,7 @@ theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h -@[deprecated "Deprecated without replacement." (since := "2024-06-09")] +@[deprecated "No deprecation message was provided." (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 1c082292d1c8c..7ddd95b920b4e 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -391,13 +391,13 @@ lemma ne_one_iff {χ : MulChar R R'} : χ ≠ 1 ↔ ∃ a : Rˣ, χ a ≠ 1 := b simp only [Ne, eq_one_iff, not_forall] /-- A multiplicative character is *nontrivial* if it takes a value `≠ 1` on a unit. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] def IsNontrivial (χ : MulChar R R') : Prop := ∃ a : Rˣ, χ a ≠ 1 set_option linter.deprecated false in /-- A multiplicative character is nontrivial iff it is not the trivial character. -/ -@[deprecated "Deprecated without replacement." (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] theorem isNontrivial_iff (χ : MulChar R R') : χ.IsNontrivial ↔ χ ≠ 1 := by simp only [IsNontrivial, Ne, MulChar.ext_iff, not_forall, one_apply_coe] @@ -573,7 +573,7 @@ theorem sum_eq_zero_of_ne_one [IsDomain R'] {χ : MulChar R R'} (hχ : χ ≠ 1) simpa only [Finset.mul_sum, ← map_mul] using b.mulLeft_bijective.sum_comp _ set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] lemma IsNontrivial.sum_eq_zero [IsDomain R'] {χ : MulChar R R'} (hχ : χ.IsNontrivial) : ∑ a, χ a = 0 := sum_eq_zero_of_ne_one ((isNontrivial_iff _).mp hχ) diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index dc90cc7b0adbe..b97cdc2c979d6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -206,11 +206,11 @@ theorem Equivalence.transitive (h : Equivalence r) : Transitive r := variable {β : Sort*} (r : β → β → Prop) (f : α → β) -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ @@ -276,7 +276,7 @@ lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_n lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a | _a, _b, hab => lt_iff_le_not_le.mp hab @@ -502,7 +502,7 @@ namespace Nat /-! Deprecated properties of inequality on `Nat` -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 2774cd126ed34..574fd2355523b 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -731,7 +731,7 @@ theorem apply_mono {f g : α →𝒄 β} {x y : α} (h₁ : f ≤ g) (h₂ : x OrderHom.apply_mono (show (f : α →o β) ≤ g from h₁) h₂ set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem ite_continuous' {p : Prop} [hp : Decidable p] (f g : α → β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => if p then f x else g x := by split_ifs <;> simp [*] diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index 6fb587d189e2c..74cbe59863ec0 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -88,7 +88,7 @@ theorem IsStrictOrder.swap (r) [IsStrictOrder α r] : IsStrictOrder α (swap r) theorem IsPartialOrder.swap (r) [IsPartialOrder α r] : IsPartialOrder α (swap r) := { @IsPreorder.swap α r _, @IsAntisymm.swap α r _ with } -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsLinearOrder.swap (r) [IsLinearOrder α r] : IsLinearOrder α (swap r) := { @IsPartialOrder.swap α r _, @IsTotal.swap α r _ with } @@ -214,7 +214,7 @@ instance (priority := 100) isStrictOrderConnected_of_isStrictTotalOrder [IsStric fun o ↦ o.elim (fun e ↦ e ▸ h) fun h' ↦ _root_.trans h' h⟩ -- see Note [lower instance priority] -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) isStrictTotalOrder_of_isStrictTotalOrder [IsStrictTotalOrder α r] : IsStrictWeakOrder α r := { isStrictWeakOrder_of_isOrderConnected with } @@ -782,7 +782,7 @@ instance [LinearOrder α] : IsStrictTotalOrder α (· < ·) where instance [LinearOrder α] : IsOrderConnected α (· < ·) := by infer_instance -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsStrictWeakOrder α (· < ·) := by infer_instance theorem transitive_le [Preorder α] : Transitive (@LE.le α _) := diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index 43f94a2400a2d..6e9f34ffd5705 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -541,7 +541,7 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm -@[deprecated "Deprecated without replacement." (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl @@ -555,7 +555,7 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' -@[deprecated "Deprecated without replacement." (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl @@ -571,11 +571,11 @@ theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max -@[deprecated "Deprecated without replacement." (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c -@[deprecated "Deprecated without replacement." (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o @@ -608,7 +608,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias aleph'_omega := aleph'_omega0 /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @@ -637,7 +637,7 @@ theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := -- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. /-- Ordinals that are cardinals are unbounded. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -645,16 +645,16 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ -@[deprecated "Deprecated without replacement." (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ /-- Infinite ordinals that are cardinals are unbounded. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded -@[deprecated "Deprecated without replacement." (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index eab39bcb9515e..ccce7048cd3df 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -147,7 +147,7 @@ protected theorem eq : #α = #β ↔ Nonempty (α ≃ β) := Quotient.eq' /-- Avoid using `Quotient.mk` to construct a `Cardinal` directly -/ -@[deprecated "Deprecated without replacement." (since := "2024-10-24")] +@[deprecated "No deprecation message was provided." (since := "2024-10-24")] theorem mk'_def (α : Type u) : @Eq Cardinal ⟦α⟧ #α := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index b289d46418d33..480fedc165b31 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -114,7 +114,7 @@ def StrictOrder.cof (r : α → α → Prop) : Cardinal := Order.cof (swap rᶜ) /-- The set in the definition of `Order.StrictOrder.cof` is nonempty. -/ -@[deprecated "Deprecated without replacement." (since := "2024-10-22")] +@[deprecated "No deprecation message was provided." (since := "2024-10-22")] theorem StrictOrder.cof_nonempty (r : α → α → Prop) [IsIrrefl α r] : { c | ∃ S : Set α, Unbounded r S ∧ #S = c }.Nonempty := @Order.cof_nonempty α _ (IsRefl.swap rᶜ) diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index efff5957a1330..cf5a7e8f4cf97 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -35,7 +35,7 @@ noncomputable def toPGame (o : Ordinal.{u}) : PGame.{u} := termination_by o decreasing_by exact ((enumIsoToType o).symm x).prop -@[deprecated "Deprecated without replacement." (since := "2024-09-22")] +@[deprecated "No deprecation message was provided." (since := "2024-09-22")] theorem toPGame_def (o : Ordinal) : o.toPGame = ⟨o.toType, PEmpty, fun x => ((enumIsoToType o).symm x).val.toPGame, PEmpty.elim⟩ := by rw [toPGame] diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index ffc9423ec3d0f..fe631ae88c5f5 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -605,14 +605,14 @@ theorem one_add_omega0 : 1 + ω = ω := by cases a <;> cases b <;> intro H <;> cases' H with _ _ H _ _ H <;> [exact H.elim; exact Nat.succ_pos _; exact Nat.succ_lt_succ H] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_omega := one_add_omega0 @[simp] theorem one_add_of_omega0_le {o} (h : ω ≤ o) : 1 + o = o := by rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, one_add_omega0] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_of_omega_le := one_add_of_omega0_le /-! ### Multiplication of ordinals -/ @@ -1185,7 +1185,7 @@ def sup {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal.{max u v} := iSup f set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sSup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sSup (Set.range f) = sup.{_, v} f := rfl @@ -1252,14 +1252,14 @@ protected theorem lt_iSup_iff {ι} {f : ι → Ordinal.{u}} {a : Ordinal.{u}} [S a < iSup f ↔ ∃ i, a < f i := lt_ciSup_iff' (bddAbove_of_small _) -@[deprecated "Deprecated without replacement." (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff +@[deprecated "No deprecation message was provided." (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff set_option linter.deprecated false in @[deprecated Ordinal.lt_iSup (since := "2024-08-27")] theorem lt_sup {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < sup.{_, v} f ↔ ∃ i, a < f i := by simpa only [not_forall, not_le] using not_congr (@sup_le_iff.{_, v} _ f a) -@[deprecated "Deprecated without replacement." (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem ne_iSup_iff_lt_iSup {ι : Type u} {f : ι → Ordinal.{max u v}} : (∀ i, f i ≠ iSup f) ↔ ∀ i, f i < iSup f := forall_congr' fun i => (Ordinal.le_iSup f i).lt_iff_ne.symm @@ -1377,7 +1377,7 @@ theorem unbounded_range_of_sup_ge {α β : Type u} (r : α → α → Prop) [IsW unbounded_range_of_le_iSup r f h set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem le_sup_shrink_equiv {s : Set Ordinal.{u}} (hs : Small.{u} s) (a) (ha : a ∈ s) : a ≤ sup.{u, u} fun x => ((@equivShrink s hs).symm x).val := by convert le_sup.{u, u} (fun x => ((@equivShrink s hs).symm x).val) ((@equivShrink s hs) ⟨a, ha⟩) @@ -1422,7 +1422,7 @@ theorem IsNormal.apply_of_isLimit {f : Ordinal.{u} → Ordinal.{v}} (H : IsNorma rw [← H.map_iSup, ho.iSup_Iio] set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_eq_sSup {s : Set Ordinal.{u}} (hs : Small.{u} s) : (sup.{u, u} fun x => (@equivShrink s hs).symm x) = sSup s := let hs' := bddAbove_iff_small.2 hs @@ -1986,12 +1986,12 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : Deprecated. If you need this value explicitly, write it in terms of `iSup`. If you just want an upper bound for the image of `op`, use that `Iio a ×ˢ Iio b` is a small set. -/ -@[deprecated "Deprecated without replacement." (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) : Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) -@[deprecated "Deprecated without replacement." (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} (ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by @@ -2012,34 +2012,34 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by refine mex_le_of_ne fun i hi => ?_ @@ -2072,18 +2072,18 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by convert (config := {transparency := .default}) @@ -2091,23 +2091,23 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := @@ -2165,7 +2165,7 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] rfl -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias one_add_nat_cast := one_add_natCast -- See note [no_index around OfNat.ofNat] @@ -2179,43 +2179,43 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n | 0 => by simp | n + 1 => by rw [Nat.mul_succ, Nat.cast_add, natCast_mul m n, Nat.cast_succ, mul_add_one] -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mul := natCast_mul @[deprecated Nat.cast_le (since := "2024-10-17")] theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_le := natCast_le @[deprecated Nat.cast_inj (since := "2024-10-17")] theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_inj := natCast_inj @[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_lt := natCast_lt @[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero @[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero @[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_pos := natCast_pos @[simp, norm_cast] @@ -2226,7 +2226,7 @@ theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by · apply (add_left_cancel n).1 rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @[simp, norm_cast] @@ -2241,7 +2241,7 @@ theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_div := natCast_div @[simp, norm_cast] @@ -2249,7 +2249,7 @@ theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by rw [← add_left_cancel, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mod := natCast_mod @[simp] @@ -2257,7 +2257,7 @@ theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n | 0 => by simp | n + 1 => by simp [lift_natCast n] -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias lift_nat_cast := lift_natCast -- See note [no_index around OfNat.ofNat] @@ -2292,13 +2292,13 @@ theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : theorem lt_omega0 {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by simp_rw [← Cardinal.ord_aleph0, Cardinal.lt_ord, lt_aleph0, card_eq_nat] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias lt_omega := lt_omega0 theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := lt_omega0.2 ⟨_, rfl⟩ -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nat_lt_omega := nat_lt_omega0 theorem omega0_pos : 0 < ω := @@ -2307,12 +2307,12 @@ theorem omega0_pos : 0 < ω := theorem omega0_ne_zero : ω ≠ 0 := omega0_pos.ne' -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_ne_zero := omega0_ne_zero theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega0 1 -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 theorem isLimit_omega0 : IsLimit ω := @@ -2320,10 +2320,10 @@ theorem isLimit_omega0 : IsLimit ω := let ⟨n, e⟩ := lt_omega0.1 h rw [e]; exact nat_lt_omega0 (n + 1)⟩ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_isLimit := isLimit_omega0 theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := @@ -2332,7 +2332,7 @@ theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := let ⟨n, e⟩ := lt_omega0.1 h rw [e, ← succ_le_iff]; exact H (n + 1)⟩ -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le := omega0_le @[simp] @@ -2344,7 +2344,7 @@ set_option linter.deprecated false in theorem sup_natCast : sup Nat.cast = ω := iSup_natCast -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias sup_nat_cast := sup_natCast theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o @@ -2354,7 +2354,7 @@ theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le_of_isLimit := omega0_le_of_isLimit theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by @@ -2372,7 +2372,7 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ intro e simp only [e, mul_zero] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias isLimit_iff_omega_dvd := isLimit_iff_omega0_dvd theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) @@ -2416,7 +2416,7 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias IsNormal.apply_omega := IsNormal.apply_omega0 @[simp] @@ -2460,7 +2460,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by · rw [ord_aleph0] exact Ordinal.isLimit_omega0 -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 649cfe8761545..ec9a34419fbf4 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -387,7 +387,7 @@ def typein (r : α → α → Prop) [IsWellOrder α r] : @PrincipalSeg α Ordina alias typein.principalSeg := typein set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : (typein.principalSeg r : α → Ordinal) = typein r := rfl @@ -513,7 +513,7 @@ noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where right_inv _ := enum_typein _ _ map_rel_iff' := enum_le_enum' _ -@[deprecated "Deprecated without replacement." (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias enumIsoOut := enumIsoToType instance small_Iio (o : Ordinal.{u}) : Small.{u} (Iio o) := @@ -925,7 +925,7 @@ theorem card_succ (o : Ordinal) : card (succ o) = card o + 1 := by theorem natCast_succ (n : ℕ) : ↑n.succ = succ (n : Ordinal) := rfl -@[deprecated "Deprecated without replacement." (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_succ := natCast_succ instance uniqueIioOne : Unique (Iio (1 : Ordinal)) where @@ -1213,7 +1213,7 @@ alias card_typein_out_lt := card_typein_toType_lt theorem mk_Iio_ord_toType {c : Cardinal} (i : c.ord.toType) : #(Iio i) < c := card_typein_toType_lt c i -@[deprecated "Deprecated without replacement." (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias mk_Iio_ord_out_α := mk_Iio_ord_toType theorem ord_injective : Injective ord := by diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean index 04236bb473897..e8b3bb9314bc5 100644 --- a/Mathlib/SetTheory/Ordinal/Enum.lean +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -33,7 +33,7 @@ termination_by o variable {s : Set Ordinal.{u}} -@[deprecated "Deprecated without replacement." (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem enumOrd_def (o : Ordinal.{u}) : enumOrd s o = sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) := by rw [enumOrd] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 4152802449ec8..84f2f8db36b0d 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -52,7 +52,7 @@ least `a`, and `Ordinal.nfpFamily_le_fp` shows this is the least ordinal with th def nfpFamily (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : Ordinal := ⨆ i, List.foldr f a i -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpFamily_eq_sup (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : nfpFamily f a = ⨆ i, List.foldr f a i := rfl @@ -238,48 +238,48 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := ⟨fun h => by @@ -287,19 +287,19 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by rw [← not_iff_not] push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -307,7 +307,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by refine ⟨fun h => ?_, fun h i hi => ?_⟩ @@ -316,13 +316,13 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : ¬ BddAbove (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [not_bddAbove_iff] @@ -347,12 +347,12 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := isNormal_derivFamily _ @@ -360,7 +360,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -368,7 +368,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by unfold derivBFamily @@ -378,7 +378,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by rw [← le_iff_derivBFamily H] @@ -387,7 +387,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : exact h i hi /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ -@[deprecated "Deprecated without replacement." (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : derivBFamily.{u, v} o f = enumOrd (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_bfamily H)] @@ -428,7 +428,7 @@ theorem iSup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) exact Ordinal.le_iSup _ _ set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : (sup fun n : ℕ => f^[n] a) = nfp f a := by refine le_antisymm ?_ (sup_le fun l => ?_) @@ -580,7 +580,7 @@ theorem nfp_add_eq_mul_omega0 {a b} (hba : b ≤ a * ω) : nfp (a + ·) b = a * exact nfp_monotone (isNormal_add_right a).monotone (Ordinal.zero_le b) · dsimp; rw [← mul_one_add, one_add_omega0] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_add_eq_mul_omega := nfp_add_eq_mul_omega0 theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω ≤ b := by @@ -593,14 +593,14 @@ theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω nth_rw 1 [← this] rwa [← add_assoc, ← mul_one_add, one_add_omega0] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias add_eq_right_iff_mul_omega_le := add_eq_right_iff_mul_omega0_le theorem add_le_right_iff_mul_omega0_le {a b : Ordinal} : a + b ≤ b ↔ a * ω ≤ b := by rw [← add_eq_right_iff_mul_omega0_le] exact (isNormal_add_right a).le_iff_eq -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias add_le_right_iff_mul_omega_le := add_le_right_iff_mul_omega0_le theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * ω + b := by @@ -612,7 +612,7 @@ theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * · rw [deriv_succ, h, add_succ] exact nfp_eq_self (add_eq_right_iff_mul_omega0_le.2 ((le_add_right _ _).trans (le_succ _))) -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias deriv_add_eq_mul_omega_add := deriv_add_eq_mul_omega0_add /-! ### Fixed points of multiplication -/ @@ -645,7 +645,7 @@ theorem nfp_mul_eq_opow_omega0 {a b : Ordinal} (hb : 0 < b) (hba : b ≤ a ^ ω) rw [← nfp_mul_one ha] exact nfp_monotone (isNormal_mul_right ha).monotone (one_le_iff_pos.2 hb) -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_mul_eq_opow_omega := nfp_mul_eq_opow_omega0 theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = b) : @@ -659,7 +659,7 @@ theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = rw [← Ne, ← one_le_iff_ne_zero] at hb exact nfp_le_fp (isNormal_mul_right ha).monotone hb (le_of_eq hab) -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias eq_zero_or_opow_omega_le_of_mul_eq_right := eq_zero_or_opow_omega0_le_of_mul_eq_right theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω ∣ b := by @@ -677,7 +677,7 @@ theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω cases' h with c hc rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega0] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias mul_eq_right_iff_opow_omega_dvd := mul_eq_right_iff_opow_omega0_dvd theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : @@ -685,7 +685,7 @@ theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : rw [← mul_eq_right_iff_opow_omega0_dvd] exact (isNormal_mul_right ha).le_iff_eq -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias mul_le_right_iff_opow_omega_dvd := mul_le_right_iff_opow_omega0_dvd theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) @@ -705,7 +705,7 @@ theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) rw [add_zero, mul_lt_mul_iff_left (opow_pos ω ha)] at this rwa [succ_le_iff] -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_mul_opow_omega_add := nfp_mul_opow_omega0_add theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : @@ -718,7 +718,7 @@ theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : · rw [deriv_succ, h] exact nfp_mul_opow_omega0_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) -@[deprecated "Deprecated without replacement." (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias deriv_mul_eq_opow_omega_mul := deriv_mul_eq_opow_omega0_mul end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5c0a4122cd004..194abeabc7c50 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -129,7 +129,7 @@ theorem not_bddAbove_principal (op : Ordinal → Ordinal → Ordinal) : exact ((le_nfp _ _).trans (ha (principal_nfp_iSup op (succ a)))).not_lt (lt_succ a) set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : Principal op (nfp (fun o' => blsub₂.{u, u, u} o' o' (@fun a _ b _ => op a b)) o) := by intro a b ha hb @@ -147,7 +147,7 @@ theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordin exact lt_blsub₂ (@fun a _ b _ => op a b) hm (hn.trans_le h) set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : Set.Unbounded (· < ·) { o | Principal op o } := fun o => ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index 724bc0096f422..a2b3354096bf0 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -447,12 +447,12 @@ set_option linter.deprecated false /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet.{u} n → Prop | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl @@ -460,46 +460,46 @@ theorem Arity.equiv_const {a : PSet.{u}} : /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 -@[deprecated "Deprecated without replacement." (since := "2024-09-02"), refl] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c | 0, _, _, _, hab, hcb => PSet.Equiv.euc hab hcb | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) -@[deprecated "Deprecated without replacement." (since := "2024-09-02"), symm] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc -@[deprecated "Deprecated without replacement." (since := "2024-09-02"), trans] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -611,7 +611,7 @@ set_option linter.deprecated false namespace Resp /-- Helper function for `PSet.eval`. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def evalAux : ∀ {n}, { f : Resp n → OfArity ZFSet.{u} ZFSet.{u} n // ∀ a b : Resp n, Resp.Equiv a b → f a = f b } @@ -626,11 +626,11 @@ def evalAux : evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := rfl @@ -640,25 +640,25 @@ end Resp /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) | mk (f) : Definable n (Resp.eval n f) -attribute [deprecated "Deprecated without replacement." (since := "2024-09-02"), instance] +attribute [deprecated "No deprecation message was provided." (since := "2024-09-02"), instance] Definable.mk /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ /-- Turns a definable function into a function that respects equivalence. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s | _, ⟨_⟩ => rfl @@ -671,7 +671,7 @@ open PSet ZFSet set_option linter.deprecated false in /-- All functions are classically definable. -/ -@[deprecated "Deprecated without replacement." (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] noncomputable def allDefinable : ∀ {n} (F : OfArity ZFSet ZFSet n), Definable n F | 0, F => let p := @Quotient.exists_rep PSet _ F @@ -707,7 +707,7 @@ theorem exact {x y : PSet} : mk x = mk y → PSet.Equiv x y := Quotient.exact set_option linter.deprecated false in -@[deprecated "Deprecated without replacement." (since := "2024-09-02"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), simp] theorem eval_mk {n f x} : (@Resp.eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) (mk x) = Resp.eval n (Resp.f f x) := rfl diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index dd9e33f4aa641..44b6433923f59 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -266,7 +266,7 @@ theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff -@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-06"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-06"))] alias smul_mem_nhds := smul_mem_nhds_smul @[to_additive (attr := simp)] diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean index ee71b9f91af12..e46af523bf2ca 100644 --- a/Mathlib/Topology/Algebra/Group/Quotient.lean +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -76,7 +76,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := QuotientGroup.isOpenQuotientMap_mk.locallyCompactSpace -@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-10-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-10-05"))] theorem continuous_smul₁ (x : G ⧸ N) : Continuous fun g : G => g • x := continuous_id.smul continuous_const @@ -101,7 +101,7 @@ instance instSecondCountableTopology [SecondCountableTopology G] : SecondCountableTopology (G ⧸ N) := ContinuousConstSMul.secondCountableTopology -@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := inferInstance @@ -117,7 +117,7 @@ instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where exact continuous_mk.comp continuous_mul continuous_inv := continuous_inv.quotient_map' _ -@[to_additive (attr := deprecated "Deprecated without replacement." (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := instTopologicalGroup N diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 091dcd05966e4..6c20290bf4fbc 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -82,7 +82,7 @@ theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : simpa [isClosed_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated "Deprecated without replacement." (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isClosedMap (s : Set β) (H : IsClosedMap f) : IsClosedMap (s.restrictPreimage f) := H.restrictPreimage s @@ -94,7 +94,7 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : simpa [isOpen_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated "Deprecated without replacement." (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index 373e5d313be0f..93c053adc16d7 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,7 +102,7 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated "Deprecated without replacement." (since := "2024-02-23")] +@[deprecated "No deprecation message was provided." (since := "2024-02-23")] theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance /-- A countable product of Polish spaces is Polish. -/ diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index be41137f7a7b4..931e8748cbfce 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -98,7 +98,7 @@ theorem noetherianSpace_iff_isCompact : NoetherianSpace α ↔ ∀ s : Set α, I instance [NoetherianSpace α] : WellFoundedLT (Closeds α) := Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› -@[deprecated "Deprecated without replacement." (since := "2024-10-07")] +@[deprecated "No deprecation message was provided." (since := "2024-10-07")] theorem NoetherianSpace.wellFounded_closeds [NoetherianSpace α] : WellFounded fun s t : Closeds α => s < t := wellFounded_lt diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 6a0a5700cbe59..2b72ca6f28b04 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -341,7 +341,7 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated "Deprecated without replacement." (since := "2024-02-15")] +@[deprecated "No deprecation message was provided." (since := "2024-02-15")] lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a export ClosedIciTopology (isClosed_ge') diff --git a/lake-manifest.json b/lake-manifest.json index aafb1ceadefe5..9e6f5223af163 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "fcc096fdce0fedca76f50f8e9d6add4eb5034962", + "rev": "994cc7f9945a217d8d115054aba1a6aab90f6d3b", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6112", From 7e3651501ccaac21324762fb104e0f12b7143559 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Sat, 23 Nov 2024 23:31:35 +0000 Subject: [PATCH 093/250] feat: more `bound` lemmas (#19285) From LeanAPAP --- Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean index bbd7e7951179d..488d67de45348 100644 --- a/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean +++ b/Mathlib/Algebra/Order/GroupWithZero/Unbundled.lean @@ -983,6 +983,11 @@ lemma mul_lt_one_of_nonneg_of_lt_one_right [MulPosMono M₀] (ha : a ≤ 1) (hb section variable [ZeroLEOneClass M₀] [PosMulMono M₀] [MulPosMono M₀] +@[bound] +protected lemma Bound.one_lt_mul : 1 ≤ a ∧ 1 < b ∨ 1 < a ∧ 1 ≤ b → 1 < a * b := by + rintro (⟨ha, hb⟩ | ⟨ha, hb⟩); exacts [one_lt_mul ha hb, one_lt_mul_of_lt_of_le ha hb] + +@[bound] lemma mul_le_one₀ (ha : a ≤ 1) (hb₀ : 0 ≤ b) (hb : b ≤ 1) : a * b ≤ 1 := one_mul (1 : M₀) ▸ mul_le_mul ha hb hb₀ zero_le_one @@ -1358,6 +1363,8 @@ lemma inv_le_iff_one_le_mul₀' (ha : 0 < a) : a⁻¹ ≤ b ↔ 1 ≤ a * b := b lemma one_le_inv₀ (ha : 0 < a) : 1 ≤ a⁻¹ ↔ a ≤ 1 := by simpa using one_le_inv_mul₀ ha (b := 1) lemma inv_le_one₀ (ha : 0 < a) : a⁻¹ ≤ 1 ↔ 1 ≤ a := by simpa using inv_mul_le_one₀ ha (b := 1) +@[bound] alias ⟨_, Bound.one_le_inv₀⟩ := one_le_inv₀ + @[bound] lemma inv_le_one_of_one_le₀ (ha : 1 ≤ a) : a⁻¹ ≤ 1 := (inv_le_one₀ <| zero_lt_one.trans_le ha).2 ha From 673d91c9789719d8b7e9545a643d8db13ad596ca Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sat, 23 Nov 2024 23:55:09 +0000 Subject: [PATCH 094/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6189 --- 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 c079482804514..67a07f2e7db05 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4a0cc9424af7acd48ce2e156e1257b718ccb79cf", + "rev": "0e076cbda9033e5cde14aa444346c6a191f33972", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6189", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..ccad69ae81b52 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-6189" 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 fdc71e836bafa..0c6fd55c26ef4 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-23 +leanprover/lean4-pr-releases:pr-release-6189 From dbdaabac1107d083fa3d1dcd6b4b512302760298 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 24 Nov 2024 00:04:58 +0000 Subject: [PATCH 095/250] chore: adaptations for nightly-2024-11-23 (#19423) Co-authored-by: leanprover-community-mathlib4-bot Co-authored-by: Kyle Miller Co-authored-by: JovanGerb Co-authored-by: Johan Commelin --- Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean | 2 +- Mathlib/Topology/Gluing.lean | 2 +- MathlibTest/Simps.lean | 2 +- lake-manifest.json | 2 +- lean-toolchain | 2 +- 5 files changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean index 5c87be9ceb0bc..df614fe6a7795 100644 --- a/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean +++ b/Mathlib/Geometry/RingedSpace/PresheafedSpace/Gluing.lean @@ -92,7 +92,7 @@ that the `U i`'s are open subspaces of the glued space. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] -structure GlueData extends GlueData (PresheafedSpace.{u, v, v} C) where +structure GlueData extends CategoryTheory.GlueData (PresheafedSpace.{u, v, v} C) where f_open : ∀ i j, IsOpenImmersion (f i j) attribute [instance] GlueData.f_open diff --git a/Mathlib/Topology/Gluing.lean b/Mathlib/Topology/Gluing.lean index c5385f98801ee..6c4f75721330d 100644 --- a/Mathlib/Topology/Gluing.lean +++ b/Mathlib/Topology/Gluing.lean @@ -83,7 +83,7 @@ Most of the times it would be easier to use the constructor `TopCat.GlueData.mk' conditions are stated in a less categorical way. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): removed @[nolint has_nonempty_instance] -structure GlueData extends GlueData TopCat where +structure GlueData extends CategoryTheory.GlueData TopCat where f_open : ∀ i j, IsOpenEmbedding (f i j) f_mono i j := (TopCat.mono_iff_injective _).mpr (f_open i j).isEmbedding.injective diff --git a/MathlibTest/Simps.lean b/MathlibTest/Simps.lean index 8def1aa96d6ef..2c3e0d2f69483 100644 --- a/MathlibTest/Simps.lean +++ b/MathlibTest/Simps.lean @@ -1203,7 +1203,7 @@ structure Foo where @[simps field field_2 field_9_fst] def myFoo : Foo := ⟨1, ⟨1, 1⟩, 1⟩ -structure Prod (X Y : Type _) extends Prod X Y +structure Prod (X Y : Type _) extends _root_.Prod X Y structure Prod2 (X Y : Type _) extends Prod X Y diff --git a/lake-manifest.json b/lake-manifest.json index ec9d70d4be933..c079482804514 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4efc0bf8d39cd7fb92e59f44ad38091ca14de2ce", + "rev": "4a0cc9424af7acd48ce2e156e1257b718ccb79cf", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", diff --git a/lean-toolchain b/lean-toolchain index ae2a1d4650fc0..fdc71e836bafa 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-22 +leanprover/lean4:nightly-2024-11-23 From 13f8b501146167f8c599e41a7f9b49054a16f6a7 Mon Sep 17 00:00:00 2001 From: leanprover-community-bot-assistant Date: Sun, 24 Nov 2024 00:16:54 +0000 Subject: [PATCH 096/250] chore(scripts): update nolints.json (#19428) I am happy to remove some nolints for you! --- scripts/nolints.json | 13 ------------- 1 file changed, 13 deletions(-) diff --git a/scripts/nolints.json b/scripts/nolints.json index a64eacb27c056..a698a174e1354 100644 --- a/scripts/nolints.json +++ b/scripts/nolints.json @@ -575,20 +575,7 @@ ["docBlame", "CategoryTheory.Limits.MultispanIndex.right"], ["docBlame", "CategoryTheory.Limits.MultispanIndex.snd"], ["docBlame", "CategoryTheory.Limits.MultispanIndex.sndFrom"], - ["docBlame", "CategoryTheory.Limits.PreservesColimit.preserves"], - ["docBlame", "CategoryTheory.Limits.PreservesLimit.preserves"], - ["docBlame", "CategoryTheory.Limits.ReflectsColimit.reflects"], - ["docBlame", "CategoryTheory.Limits.ReflectsColimitsOfShape.reflectsColimit"], - ["docBlame", - "CategoryTheory.Limits.ReflectsColimitsOfSize.reflectsColimitsOfShape"], - ["docBlame", "CategoryTheory.Limits.ReflectsLimit.reflects"], - ["docBlame", "CategoryTheory.Limits.ReflectsLimitsOfShape.reflectsLimit"], - ["docBlame", - "CategoryTheory.Limits.ReflectsLimitsOfSize.reflectsLimitsOfShape"], ["docBlame", "CategoryTheory.Monad.CreatesColimitOfIsSplitPair.out"], - ["docBlame", "CategoryTheory.Monad.PreservesColimitOfIsReflexivePair.out"], - ["docBlame", "CategoryTheory.Monad.PreservesColimitOfIsSplitPair.out"], - ["docBlame", "CategoryTheory.Monad.ReflectsColimitOfIsSplitPair.out"], ["docBlame", "CategoryTheory.Pretriangulated.Triangle.homMk"], ["docBlame", "CategoryTheory.Pretriangulated.Triangle.isoMk"], ["docBlame", "CategoryTheory.Triangulated.Octahedron.m₁"], From 54647d73257afdbff4d2a932c578d1e138b99e39 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Sat, 23 Nov 2024 17:36:35 -0800 Subject: [PATCH 097/250] fixes --- Mathlib/Algebra/Quaternion.lean | 4 +--- Mathlib/Topology/OmegaCompletePartialOrder.lean | 2 +- 2 files changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 69ce497aff36a..8cac9eec86da2 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -744,8 +744,6 @@ namespace Quaternion variable {S T R : Type*} [CommRing R] (r x y z : R) (a b c : ℍ[R]) -export QuaternionAlgebra (re imI imJ imK) - /-- Coercion `R → ℍ[R]`. -/ @[coe] def coe : R → ℍ[R] := QuaternionAlgebra.coe @@ -1191,7 +1189,7 @@ theorem normSq_le_zero : normSq a ≤ 0 ↔ a = 0 := normSq_nonneg.le_iff_eq.trans normSq_eq_zero instance instNontrivial : Nontrivial ℍ[R] where - exists_pair_ne := ⟨0, 1, mt (congr_arg re) zero_ne_one⟩ + exists_pair_ne := ⟨0, 1, mt (congr_arg QuaternionAlgebra.re) zero_ne_one⟩ instance : NoZeroDivisors ℍ[R] where eq_zero_or_eq_zero_of_mul_eq_zero {a b} hab := diff --git a/Mathlib/Topology/OmegaCompletePartialOrder.lean b/Mathlib/Topology/OmegaCompletePartialOrder.lean index c7890d3939837..cce0eb4e71f4a 100644 --- a/Mathlib/Topology/OmegaCompletePartialOrder.lean +++ b/Mathlib/Topology/OmegaCompletePartialOrder.lean @@ -106,7 +106,7 @@ theorem notBelow_isOpen : IsOpen (notBelow y) := by end notBelow -open Scott hiding IsOpen +open Scott hiding IsOpen IsOpen.isUpperSet open OmegaCompletePartialOrder From 1effc5b505efb1fd29ad88b864d231662ab2e40f Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sun, 24 Nov 2024 06:37:53 +0000 Subject: [PATCH 098/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6194 --- 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 c079482804514..39f4f277c45e5 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4a0cc9424af7acd48ce2e156e1257b718ccb79cf", + "rev": "3edeb1e515695b57b20d8c128db6be3703cece8c", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6194", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..b2f3ff65f4d9d 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-6194" 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 fdc71e836bafa..7b38e2425baa4 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-23 +leanprover/lean4-pr-releases:pr-release-6194 From 9fc2f4be35a92ff1f0ec4ea21895bdebf4087d27 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 24 Nov 2024 18:19:38 +1100 Subject: [PATCH 099/250] fixes for leanprover/lean4#6194 --- .../Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean | 4 ++-- MathlibTest/Expr.lean | 2 +- lake-manifest.json | 2 +- 3 files changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean index 8a5a9488ae1a9..2f6eb09ce7de1 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/Datatypes.lean @@ -69,7 +69,7 @@ instance : UsableInSimplexAlgorithm DenseMatrix where for ⟨i, j, v⟩ in vals do data := data.modify i fun row => row.set! j v return ⟨data⟩ - swapRows mat i j := ⟨mat.data.swap! i j⟩ + swapRows mat i j := ⟨mat.data.swapIfInBounds i j⟩ subtractRow mat i j coef := let newData : Array (Array Rat) := mat.data.modify j fun row => row.zipWith mat.data[i]! fun x y => x - coef * y @@ -101,7 +101,7 @@ instance : UsableInSimplexAlgorithm SparseMatrix where if v != 0 then data := data.modify i fun row => row.insert j v return ⟨data⟩ - swapRows mat i j := ⟨mat.data.swap! i j⟩ + swapRows mat i j := ⟨mat.data.swapIfInBounds i j⟩ subtractRow mat i j coef := let newData := mat.data.modify j fun row => mat.data[i]!.fold (fun cur k val => diff --git a/MathlibTest/Expr.lean b/MathlibTest/Expr.lean index 1e4b9fd11388d..84efba4ecb2f9 100644 --- a/MathlibTest/Expr.lean +++ b/MathlibTest/Expr.lean @@ -13,7 +13,7 @@ def reorderLastArguments : Expr → Expr := Expr.replaceRec fun r e ↦ let n := e.getAppNumArgs if n ≥ 2 then - mkAppN e.getAppFn <| e.getAppArgs.map r |>.swap! (n - 1) (n - 2) + mkAppN e.getAppFn <| e.getAppArgs.map r |>.swapIfInBounds (n - 1) (n - 2) else none diff --git a/lake-manifest.json b/lake-manifest.json index 39f4f277c45e5..7bafe87540e02 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "3edeb1e515695b57b20d8c128db6be3703cece8c", + "rev": "c87de447533223d663f5fd67fd1e602caf725da2", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6194", From 3ece930d0a4a55679efa52b1a825ac93b2469a06 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Sun, 24 Nov 2024 07:24:05 +0000 Subject: [PATCH 100/250] chore: import Plausible in Tactic.Common (#19429) --- Mathlib/Tactic/Common.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Tactic/Common.lean b/Mathlib/Tactic/Common.lean index 90340cd10a34e..3c61c4a884912 100644 --- a/Mathlib/Tactic/Common.lean +++ b/Mathlib/Tactic/Common.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ --- First import Aesop and Qq +-- First import Aesop, Qq, and Plausible import Aesop import Qq +import Plausible -- Tools for analysing imports, like `#find_home`, `#minimize_imports`, ... import ImportGraph.Imports From 2a61a5fa2521ec1a34cb24de5377b2c74d958400 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sun, 24 Nov 2024 07:34:55 +0000 Subject: [PATCH 101/250] Trigger CI for https://github.com/leanprover/lean4/pull/6194 --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 7bafe87540e02..7f28c5a4c65df 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c87de447533223d663f5fd67fd1e602caf725da2", + "rev": "897bb634d59160910345ea1eed665f622d9146de", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6194", From db0487a4dbe7caa9f45514455f4edddbc56d0505 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sun, 24 Nov 2024 10:01:56 +0000 Subject: [PATCH 102/250] chore: bump to nightly-2024-11-24 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index fdc71e836bafa..22be24ab2e033 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-23 +leanprover/lean4:nightly-2024-11-24 From 83ce5b0136c9d9b856df5fc8fb4baa1c0527372f Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Sun, 24 Nov 2024 10:18:35 +0000 Subject: [PATCH 103/250] feat(CategoryTheory/Over): more API for `post` (#19312) --- Mathlib/CategoryTheory/Comma/Over.lean | 181 +++++++++++++++++++++++++ 1 file changed, 181 insertions(+) diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index e37bf7ec99866..305664eb69cad 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -161,6 +161,13 @@ theorem map_map_left : ((map f).map g).left = g.left := rfl end +/-- If `f` is an isomorphism, `map f` is an equivalence of categories. -/ +def mapIso {Y : T} (f : X ≅ Y) : Over X ≌ Over Y := + Comma.mapRightIso _ <| Discrete.natIso fun _ ↦ f + +@[simp] lemma mapIso_functor {Y : T} (f : X ≅ Y) : (mapIso f).functor = map f.hom := rfl +@[simp] lemma mapIso_inverse {Y : T} (f : X ≅ Y) : (mapIso f).inverse = map f.inv := rfl + section coherences /-! This section proves various equalities between functors that @@ -186,6 +193,7 @@ theorem mapId_eq (Y : T) : map (𝟙 Y) = 𝟭 _ := by simp /-- The natural isomorphism arising from `mapForget_eq`. -/ +@[simps!] def mapId (Y : T) : map (𝟙 Y) ≅ 𝟭 _ := eqToIso (mapId_eq Y) -- NatIso.ofComponents fun X => isoMk (Iso.refl _) @@ -215,9 +223,16 @@ theorem mapComp_eq {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : simp /-- The natural isomorphism arising from `mapComp_eq`. -/ +@[simps!] def mapComp {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) ≅ (map f) ⋙ (map g) := eqToIso (mapComp_eq f g) +/-- If `f = g`, then `map f` is naturally isomorphic to `map g`. -/ +@[simps!] +def mapCongr {X Y : T} (f g : X ⟶ Y) (h : f = g) : + map f ≅ map g := + NatIso.ofComponents (fun A ↦ eqToIso (by rw [h])) + variable (T) in /-- The functor defined by the over categories.-/ @[simps] def mapFunctor : T ⥤ Cat where @@ -317,6 +332,57 @@ def post (F : T ⥤ D) : Over X ⥤ Over (F.obj X) where map f := Over.homMk (F.map f.left) (by simp only [Functor.id_obj, mk_left, Functor.const_obj_obj, mk_hom, ← F.map_comp, w]) +lemma post_comp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) = post (X := X) F ⋙ post G := + rfl + +/-- `post (F ⋙ G)` is isomorphic (actually equal) to `post F ⋙ post G`. -/ +@[simps!] +def postComp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) ≅ post F ⋙ post G := + NatIso.ofComponents (fun X ↦ Iso.refl _) + +/-- A natural transformation `F ⟶ G` induces a natural transformation on +`Over X` up to `Under.map`. -/ +@[simps] +def postMap {F G : T ⥤ D} (e : F ⟶ G) : post F ⋙ map (e.app X) ⟶ post G where + app Y := Over.homMk (e.app Y.left) + +/-- If `F` and `G` are naturally isomorphic, then `Over.post F` and `Over.post G` are also naturally +isomorphic up to `Over.map` -/ +@[simps!] +def postCongr {F G : T ⥤ D} (e : F ≅ G) : post F ⋙ map (e.hom.app X) ≅ post G := + NatIso.ofComponents (fun A ↦ Over.isoMk (e.app A.left)) + +variable (X) (F : T ⥤ D) + +instance [F.Faithful] : (Over.post (X := X) F).Faithful where + map_injective {A B} f g h := by + ext + exact F.map_injective (congrArg CommaMorphism.left h) + +instance [F.Faithful] [F.Full] : (Over.post (X := X) F).Full where + map_surjective {A B} f := by + obtain ⟨a, ha⟩ := F.map_surjective f.left + have w : a ≫ B.hom = A.hom := F.map_injective <| by simpa [ha] using Over.w _ + exact ⟨Over.homMk a, by ext; simpa⟩ + +instance [F.Full] [F.EssSurj] : (Over.post (X := X) F).EssSurj where + mem_essImage B := by + obtain ⟨A', ⟨e⟩⟩ := Functor.EssSurj.mem_essImage (F := F) B.left + obtain ⟨f, hf⟩ := F.map_surjective (e.hom ≫ B.hom) + exact ⟨Over.mk f, ⟨Over.isoMk e⟩⟩ + +instance [F.IsEquivalence] : (Over.post (X := X) F).IsEquivalence where + +/-- An equivalence of categories induces an equivalence on over categories. -/ +@[simps] +def postEquiv (F : T ≌ D) : Over X ≌ Over (F.functor.obj X) where + functor := Over.post F.functor + inverse := Over.post (X := F.functor.obj X) F.inverse ⋙ Over.map (F.unitIso.inv.app X) + unitIso := NatIso.ofComponents (fun A ↦ Over.isoMk (F.unitIso.app A.left)) + counitIso := NatIso.ofComponents (fun A ↦ Over.isoMk (F.counitIso.app A.left)) + end Over namespace CostructuredArrow @@ -457,6 +523,13 @@ theorem map_map_right : ((map f).map g).right = g.right := rfl end +/-- If `f` is an isomorphism, `map f` is an equivalence of categories. -/ +def mapIso {Y : T} (f : X ≅ Y) : Under Y ≌ Under X := + Comma.mapLeftIso _ <| Discrete.natIso fun _ ↦ f.symm + +@[simp] lemma mapIso_functor {Y : T} (f : X ≅ Y) : (mapIso f).functor = map f.hom := rfl +@[simp] lemma mapIso_inverse {Y : T} (f : X ≅ Y) : (mapIso f).inverse = map f.inv := rfl + section coherences /-! This section proves various equalities between functors that @@ -476,6 +549,7 @@ theorem mapId_eq (Y : T) : map (𝟙 Y) = 𝟭 _ := by simp /-- Mapping by the identity morphism is just the identity functor. -/ +@[simps!] def mapId (Y : T) : map (𝟙 Y) ≅ 𝟭 _ := eqToIso (mapId_eq Y) /-- Mapping by `f` and then forgetting is the same as forgetting. -/ @@ -504,9 +578,16 @@ theorem mapComp_eq {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : simp /-- The natural isomorphism arising from `mapComp_eq`. -/ +@[simps!] def mapComp {Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : map (f ≫ g) ≅ map g ⋙ map f := eqToIso (mapComp_eq f g) +/-- If `f = g`, then `map f` is naturally isomorphic to `map g`. -/ +@[simps!] +def mapCongr {X Y : T} (f g : X ⟶ Y) (h : f = g) : + map f ≅ map g := + NatIso.ofComponents (fun A ↦ eqToIso (by rw [h])) + variable (T) in /-- The functor defined by the under categories.-/ @[simps] def mapFunctor : Tᵒᵖ ⥤ Cat where @@ -570,6 +651,58 @@ def post {X : T} (F : T ⥤ D) : Under X ⥤ Under (F.obj X) where map f := Under.homMk (F.map f.right) (by simp only [Functor.id_obj, Functor.const_obj_obj, mk_right, mk_hom, ← F.map_comp, w]) +lemma post_comp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) = post (X := X) F ⋙ post G := + rfl + +/-- `post (F ⋙ G)` is isomorphic (actually equal) to `post F ⋙ post G`. -/ +@[simps!] +def postComp {E : Type*} [Category E] (F : T ⥤ D) (G : D ⥤ E) : + post (X := X) (F ⋙ G) ≅ post F ⋙ post G := + NatIso.ofComponents (fun X ↦ Iso.refl _) + +/-- A natural transformation `F ⟶ G` induces a natural transformation on +`Under X` up to `Under.map`. -/ +@[simps] +def postMap {F G : T ⥤ D} (e : F ⟶ G) : post (X := X) F ⟶ post G ⋙ map (e.app X) where + app Y := Under.homMk (e.app Y.right) + +/-- If `F` and `G` are naturally isomorphic, then `Under.post F` and `Under.post G` are also +naturally isomorphic up to `Under.map` -/ +@[simps!] +def postCongr {F G : T ⥤ D} (e : F ≅ G) : post F ≅ post G ⋙ map (e.hom.app X) := + NatIso.ofComponents (fun A ↦ Under.isoMk (e.app A.right)) + +variable (X) (F : T ⥤ D) + +instance [F.Faithful] : (Under.post (X := X) F).Faithful where + map_injective {A B} f g h := by + ext + exact F.map_injective (congrArg CommaMorphism.right h) + +instance [F.Faithful] [F.Full] : (Under.post (X := X) F).Full where + map_surjective {A B} f := by + obtain ⟨a, ha⟩ := F.map_surjective f.right + dsimp at a + have w : A.hom ≫ a = B.hom := F.map_injective <| by simpa [ha] using Under.w f + exact ⟨Under.homMk a, by ext; simpa⟩ + +instance [F.Full] [F.EssSurj] : (Under.post (X := X) F).EssSurj where + mem_essImage B := by + obtain ⟨B', ⟨e⟩⟩ := Functor.EssSurj.mem_essImage (F := F) B.right + obtain ⟨f, hf⟩ := F.map_surjective (B.hom ≫ e.inv) + exact ⟨Under.mk f, ⟨Under.isoMk e⟩⟩ + +instance [F.IsEquivalence] : (Under.post (X := X) F).IsEquivalence where + +/-- An equivalence of categories induces an equivalence on under categories. -/ +@[simps] +def postEquiv (F : T ≌ D) : Under X ≌ Under (F.functor.obj X) where + functor := post F.functor + inverse := post (X := F.functor.obj X) F.inverse ⋙ Under.map (F.unitIso.hom.app X) + unitIso := NatIso.ofComponents (fun A ↦ Under.isoMk (F.unitIso.app A.right)) + counitIso := NatIso.ofComponents (fun A ↦ Under.isoMk (F.counitIso.app A.right)) + end Under namespace StructuredArrow @@ -774,4 +907,52 @@ def ofDiagEquivalence' (X : T × T) : end CostructuredArrow +section Opposite + +open Opposite + +variable (X : T) + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Over.opToOpUnder : Over (op X) ⥤ (Under X)ᵒᵖ where + obj Y := ⟨Under.mk Y.hom.unop⟩ + map {Z Y} f := ⟨Under.homMk (f.left.unop) (by dsimp; rw [← unop_comp, Over.w])⟩ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Under.opToOverOp : (Under X)ᵒᵖ ⥤ Over (op X) where + obj Y := Over.mk (Y.unop.hom.op) + map {Z Y} f := Over.homMk f.unop.right.op <| by dsimp; rw [← Under.w f.unop, op_comp] + +/-- `Over.opToOpUnder` is an equivalence of categories. -/ +@[simps] +def Over.opEquivOpUnder : Over (op X) ≌ (Under X)ᵒᵖ where + functor := Over.opToOpUnder X + inverse := Under.opToOverOp X + unitIso := Iso.refl _ + counitIso := Iso.refl _ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Under.opToOpOver : Under (op X) ⥤ (Over X)ᵒᵖ where + obj Y := ⟨Over.mk Y.hom.unop⟩ + map {Z Y} f := ⟨Over.homMk (f.right.unop) (by dsimp; rw [← unop_comp, Under.w])⟩ + +/-- The canonical functor by reversing structure arrows. -/ +@[simps] +def Over.opToUnderOp : (Over X)ᵒᵖ ⥤ Under (op X) where + obj Y := Under.mk (Y.unop.hom.op) + map {Z Y} f := Under.homMk f.unop.left.op <| by dsimp; rw [← Over.w f.unop, op_comp] + +/-- `Under.opToOpOver` is an equivalence of categories. -/ +@[simps] +def Under.opEquivOpOver : Under (op X) ≌ (Over X)ᵒᵖ where + functor := Under.opToOpOver X + inverse := Over.opToUnderOp X + unitIso := Iso.refl _ + counitIso := Iso.refl _ + +end Opposite + end CategoryTheory From 6c225037eb50056d12f06584a2dd4ee0008910d9 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Sun, 24 Nov 2024 13:54:52 +0000 Subject: [PATCH 104/250] chore(AlgebraicGeometry) add `Scheme.Hom.appTop` (#19130) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit to hide the defeq abuse `f ⁻¹ᵁ ⊤ = ⊤` --- Mathlib/AlgebraicGeometry/AffineScheme.lean | 52 +++---- Mathlib/AlgebraicGeometry/AffineSpace.lean | 135 ++++++++---------- .../GammaSpecAdjunction.lean | 25 ++-- .../AlgebraicGeometry/Morphisms/Affine.lean | 2 +- .../Morphisms/AffineAnd.lean | 29 ++-- .../Morphisms/ClosedImmersion.lean | 28 ++-- .../AlgebraicGeometry/Morphisms/Finite.lean | 2 +- .../AlgebraicGeometry/Morphisms/IsIso.lean | 4 +- .../Morphisms/RingHomProperties.lean | 38 ++--- .../Morphisms/SurjectiveOnStalks.lean | 1 - Mathlib/AlgebraicGeometry/Restrict.lean | 25 +++- Mathlib/AlgebraicGeometry/Scheme.lean | 33 ++++- Mathlib/AlgebraicGeometry/Stalk.lean | 9 +- .../AlgebraicGeometry/ValuativeCriterion.lean | 6 +- 14 files changed, 214 insertions(+), 175 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/AffineScheme.lean b/Mathlib/AlgebraicGeometry/AffineScheme.lean index a83addf8247e5..96002facc9d47 100644 --- a/Mathlib/AlgebraicGeometry/AffineScheme.lean +++ b/Mathlib/AlgebraicGeometry/AffineScheme.lean @@ -67,12 +67,12 @@ def Scheme.isoSpec (X : Scheme) [IsAffine X] : X ≅ Spec Γ(X, ⊤) := @[reassoc] theorem Scheme.isoSpec_hom_naturality {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : - X.isoSpec.hom ≫ Spec.map (f.app ⊤) = f ≫ Y.isoSpec.hom := by + X.isoSpec.hom ≫ Spec.map (f.appTop) = f ≫ Y.isoSpec.hom := by simp only [isoSpec, asIso_hom, Scheme.toSpecΓ_naturality] @[reassoc] theorem Scheme.isoSpec_inv_naturality {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : - Spec.map (f.app ⊤) ≫ Y.isoSpec.inv = X.isoSpec.inv ≫ f := by + Spec.map (f.appTop) ≫ Y.isoSpec.inv = X.isoSpec.inv ≫ f := by rw [Iso.eq_inv_comp, isoSpec, asIso_hom, ← Scheme.toSpecΓ_naturality_assoc, isoSpec, asIso_inv, IsIso.hom_inv_id, Category.comp_id] @@ -112,13 +112,13 @@ theorem isAffine_of_isIso {X Y : Scheme} (f : X ⟶ Y) [IsIso f] [h : IsAffine Y to the arrow of the morphism on prime spectra induced by the map on global sections. -/ noncomputable def arrowIsoSpecΓOfIsAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) : - Arrow.mk f ≅ Arrow.mk (Spec.map (Scheme.Γ.map f.op)) := + Arrow.mk f ≅ Arrow.mk (Spec.map f.appTop) := Arrow.isoMk X.isoSpec Y.isoSpec (ΓSpec.adjunction.unit_naturality _) /-- If `f : A ⟶ B` is a ring homomorphism, the corresponding arrow is isomorphic to the arrow of the morphism induced on global sections by the map on prime spectra. -/ def arrowIsoΓSpecOfIsAffine {A B : CommRingCat} (f : A ⟶ B) : - Arrow.mk f ≅ Arrow.mk ((Spec.map f).app ⊤) := + Arrow.mk f ≅ Arrow.mk ((Spec.map f).appTop) := Arrow.isoMk (Scheme.ΓSpecIso _).symm (Scheme.ΓSpecIso _).symm (Scheme.ΓSpecIso_inv_naturality f).symm @@ -134,7 +134,7 @@ theorem Scheme.isoSpec_Spec (R : CommRingCat.{u}) : (Spec R).isoSpec.inv = Spec.map (Scheme.ΓSpecIso R).inv := congr($(isoSpec_Spec R).inv) -lemma ext_of_isAffine {X Y : Scheme} [IsAffine Y] {f g : X ⟶ Y} (e : f.app ⊤ = g.app ⊤) : +lemma ext_of_isAffine {X Y : Scheme} [IsAffine Y] {f g : X ⟶ Y} (e : f.appTop = g.appTop) : f = g := by rw [← cancel_mono Y.toSpecΓ, Scheme.toSpecΓ_naturality, Scheme.toSpecΓ_naturality, e] @@ -315,21 +315,24 @@ lemma isoSpec_hom_base_apply (x : U) : congr 1 exact IsLocalRing.comap_closedPoint (U.stalkIso x).inv -lemma isoSpec_inv_app_top : - hU.isoSpec.inv.app ⊤ = U.topIso.hom ≫ (Scheme.ΓSpecIso Γ(X, U)).inv := by +lemma isoSpec_inv_appTop : + hU.isoSpec.inv.appTop = U.topIso.hom ≫ (Scheme.ΓSpecIso Γ(X, U)).inv := by simp only [Scheme.Opens.toScheme_presheaf_obj, isoSpec_inv, Scheme.isoSpec, asIso_inv, - Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, Scheme.inv_app_top, + Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, Scheme.inv_appTop, Scheme.Opens.topIso_hom, Scheme.ΓSpecIso_inv_naturality, IsIso.inv_comp_eq] - rw [Scheme.toSpecΓ_app_top] + rw [Scheme.toSpecΓ_appTop] erw [Iso.hom_inv_id_assoc] -lemma isoSpec_hom_app_top : - hU.isoSpec.hom.app ⊤ = (Scheme.ΓSpecIso Γ(X, U)).hom ≫ U.topIso.inv := by - have := congr(inv $hU.isoSpec_inv_app_top) +lemma isoSpec_hom_appTop : + hU.isoSpec.hom.appTop = (Scheme.ΓSpecIso Γ(X, U)).hom ≫ U.topIso.inv := by + have := congr(inv $hU.isoSpec_inv_appTop) rw [IsIso.inv_comp, IsIso.Iso.inv_inv, IsIso.Iso.inv_hom] at this have := (Scheme.Γ.map_inv hU.isoSpec.inv.op).trans this rwa [← op_inv, IsIso.Iso.inv_inv] at this +@[deprecated (since := "2024-11-16")] alias isoSpec_inv_app_top := isoSpec_inv_appTop +@[deprecated (since := "2024-11-16")] alias isoSpec_hom_app_top := isoSpec_hom_appTop + /-- The open immersion `Spec Γ(X, U) ⟶ X` for an affine `U`. -/ def fromSpec : Spec Γ(X, U) ⟶ X := @@ -364,7 +367,7 @@ theorem map_fromSpec {V : X.Opens} (hV : IsAffineOpen V) (f : op U ⟶ op V) : conv_rhs => rw [fromSpec, ← X.homOfLE_ι (V := U) f.unop.le, isoSpec_inv, Category.assoc, ← Scheme.isoSpec_inv_naturality_assoc, - ← Spec.map_comp_assoc, Scheme.homOfLE_app, ← Functor.map_comp] + ← Spec.map_comp_assoc, Scheme.homOfLE_appTop, ← Functor.map_comp] rw [fromSpec, isoSpec_inv, Category.assoc, ← Spec.map_comp_assoc, ← Functor.map_comp] rfl @@ -377,21 +380,22 @@ lemma Spec_map_appLE_fromSpec (f : X ⟶ Y) {V : X.Opens} {U : Y.Opens} simp_rw [← Scheme.homOfLE_ι _ i] rw [Category.assoc, ← morphismRestrict_ι, ← Category.assoc _ (f ∣_ U) U.ι, ← @Scheme.isoSpec_inv_naturality_assoc, - ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_app, morphismRestrict_app, - Scheme.homOfLE_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, + ← Spec.map_comp_assoc, ← Spec.map_comp_assoc, Scheme.comp_appTop, morphismRestrict_appTop, + Scheme.homOfLE_appTop, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.appLE_map, Scheme.Hom.map_appLE] lemma fromSpec_top [IsAffine X] : (isAffineOpen_top X).fromSpec = X.isoSpec.inv := by - rw [fromSpec, isoSpec_inv, Category.assoc, ← @Scheme.isoSpec_inv_naturality, Scheme.Opens.ι_app, - ← Spec.map_comp_assoc, ← X.presheaf.map_comp, ← op_comp, eqToHom_comp_homOfLE, - ← eqToHom_eq_homOfLE rfl, eqToHom_refl, op_id, X.presheaf.map_id, Spec.map_id, Category.id_comp] + rw [fromSpec, isoSpec_inv, Category.assoc, ← @Scheme.isoSpec_inv_naturality, + ← Spec.map_comp_assoc, Scheme.Opens.ι_appTop, ← X.presheaf.map_comp, ← op_comp, + eqToHom_comp_homOfLE, ← eqToHom_eq_homOfLE rfl, eqToHom_refl, op_id, X.presheaf.map_id, + Spec.map_id, Category.id_comp] lemma fromSpec_app_of_le (V : X.Opens) (h : U ≤ V) : hU.fromSpec.app V = X.presheaf.map (homOfLE h).op ≫ (Scheme.ΓSpecIso Γ(X, U)).inv ≫ (Spec _).presheaf.map (homOfLE le_top).op := by have : U.ι ⁻¹ᵁ V = ⊤ := eq_top_iff.mpr fun x _ ↦ h x.2 rw [IsAffineOpen.fromSpec, Scheme.comp_app, Scheme.Opens.ι_app, Scheme.app_eq _ this, - IsAffineOpen.isoSpec_inv_app_top] + ← Scheme.Hom.appTop, IsAffineOpen.isoSpec_inv_appTop] simp only [Scheme.Opens.toScheme_presheaf_map, Scheme.Opens.topIso_hom, Category.assoc, ← X.presheaf.map_comp_assoc] rfl @@ -470,7 +474,7 @@ theorem ΓSpecIso_hom_fromSpec_app : simp only [fromSpec, Scheme.comp_coeBase, Opens.map_comp_obj, Scheme.comp_app, Scheme.Opens.ι_app_self, eqToHom_op, Scheme.app_eq _ U.ι_preimage_self, Scheme.Opens.toScheme_presheaf_map, eqToHom_unop, eqToHom_map U.ι.opensFunctor, Opens.map_top, - isoSpec_inv_app_top, Scheme.Opens.topIso_hom, Category.assoc, ← Functor.map_comp_assoc, + isoSpec_inv_appTop, Scheme.Opens.topIso_hom, Category.assoc, ← Functor.map_comp_assoc, eqToHom_trans, eqToHom_refl, X.presheaf.map_id, Category.id_comp, Iso.hom_inv_id_assoc] @[elementwise] @@ -953,10 +957,10 @@ lemma specTargetImageRingHom_surjective : Function.Surjective (specTargetImageRi Ideal.Quotient.mk_surjective lemma specTargetImageFactorization_app_injective : - Function.Injective <| (specTargetImageFactorization f).app ⊤ := by + Function.Injective <| (specTargetImageFactorization f).appTop := by let φ : A ⟶ Γ(X, ⊤) := (((ΓSpec.adjunction).homEquiv X (op A)).symm f).unop let φ' : specTargetImage f ⟶ Scheme.Γ.obj (op X) := RingHom.kerLift φ - show Function.Injective <| ((ΓSpec.adjunction.homEquiv X _) φ'.op).app ⊤ + show Function.Injective <| ((ΓSpec.adjunction.homEquiv X _) φ'.op).appTop rw [ΓSpec_adjunction_homEquiv_eq] apply (RingHom.kerLift_injective φ).comp exact ((ConcreteCategory.isIso_iff_bijective (Scheme.ΓSpecIso _).hom).mp inferInstance).injective @@ -993,7 +997,7 @@ def affineTargetImageInclusion (f : X ⟶ Y) : affineTargetImage f ⟶ Y := Spec.map (specTargetImageRingHom (f ≫ Y.isoSpec.hom)) ≫ Y.isoSpec.inv lemma affineTargetImageInclusion_app_surjective : - Function.Surjective <| (affineTargetImageInclusion f).app ⊤ := by + Function.Surjective <| (affineTargetImageInclusion f).appTop := by simp only [Scheme.comp_coeBase, Opens.map_comp_obj, Opens.map_top, Scheme.comp_app, CommRingCat.coe_comp, affineTargetImageInclusion] apply Function.Surjective.comp @@ -1013,7 +1017,7 @@ def affineTargetImageFactorization (f : X ⟶ Y) : X ⟶ affineTargetImage f := specTargetImageFactorization (f ≫ Y.isoSpec.hom) lemma affineTargetImageFactorization_app_injective : - Function.Injective <| (affineTargetImageFactorization f).app ⊤ := + Function.Injective <| (affineTargetImageFactorization f).appTop := specTargetImageFactorization_app_injective (f ≫ Y.isoSpec.hom) @[reassoc (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/AffineSpace.lean b/Mathlib/AlgebraicGeometry/AffineSpace.lean index 8e43ca277f77f..805e43c4928d2 100644 --- a/Mathlib/AlgebraicGeometry/AffineSpace.lean +++ b/Mathlib/AlgebraicGeometry/AffineSpace.lean @@ -71,7 +71,7 @@ Use `homOverEquiv` instead. -/ @[simps] def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where - toFun f i := f.app ⊤ ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) + toFun f i := f.appTop ((Scheme.ΓSpecIso ℤ[n]).inv (.X i)) invFun v := X.toSpecΓ ≫ Spec.map (MvPolynomial.eval₂Hom ((algebraMap ℤ _).comp ULift.ringEquiv.toRingHom) v) left_inv f := by @@ -80,7 +80,7 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where rw [Adjunction.homEquiv_symm_apply, Adjunction.homEquiv_symm_apply] simp only [Functor.rightOp_obj, Scheme.Γ_obj, Scheme.Spec_obj, algebraMap_int_eq, RingEquiv.toRingHom_eq_coe, TopologicalSpace.Opens.map_top, Functor.rightOp_map, op_comp, - Scheme.Γ_map, unop_comp, Quiver.Hom.unop_op, Scheme.comp_app, Scheme.toSpecΓ_app_top, + Scheme.Γ_map, unop_comp, Quiver.Hom.unop_op, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, ΓSpec.adjunction_counit_app, Category.assoc, Iso.cancel_iso_inv_left, ← Iso.eq_inv_comp] apply of_mvPolynomial_int_ext @@ -91,13 +91,13 @@ def toSpecMvPolyIntEquiv : (X ⟶ Spec ℤ[n]) ≃ (n → Γ(X, ⊤)) where ext i simp only [algebraMap_int_eq, RingEquiv.toRingHom_eq_coe, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, - Scheme.toSpecΓ_app_top, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, + Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, CommRingCat.coe_of] erw [Iso.hom_inv_id_apply] rw [coe_eval₂Hom, eval₂_X] lemma toSpecMvPolyIntEquiv_comp {X Y : Scheme} (f : X ⟶ Y) (g : Y ⟶ Spec ℤ[n]) (i) : - toSpecMvPolyIntEquiv n (f ≫ g) i = f.app ⊤ (toSpecMvPolyIntEquiv n g i) := rfl + toSpecMvPolyIntEquiv n (f ≫ g) i = f.appTop (toSpecMvPolyIntEquiv n g i) := rfl variable {n} in /-- The standard coordinates of `𝔸(n; S)`. -/ @@ -122,16 +122,16 @@ lemma homOfVector_toSpecMvPoly : homOfVector f v ≫ toSpecMvPoly n S = (toSpecMvPolyIntEquiv n).symm v := pullback.lift_snd _ _ _ -@[local simp] -lemma homOfVector_app_top_coord (i) : - (homOfVector f v).app ⊤ (coord S i) = v i := by +@[simp] +lemma homOfVector_appTop_coord (i) : + (homOfVector f v).appTop (coord S i) = v i := by rw [coord, ← toSpecMvPolyIntEquiv_comp, homOfVector_toSpecMvPoly, Equiv.apply_symm_apply] @[ext 1100] lemma hom_ext {f g : X ⟶ 𝔸(n; S)} (h₁ : f ≫ 𝔸(n; S) ↘ S = g ≫ 𝔸(n; S) ↘ S) - (h₂ : ∀ i, f.app ⊤ (coord S i) = g.app ⊤ (coord S i)) : f = g := by + (h₂ : ∀ i, f.appTop (coord S i) = g.appTop (coord S i)) : f = g := by apply pullback.hom_ext h₁ show f ≫ toSpecMvPoly _ _ = g ≫ toSpecMvPoly _ _ apply (toSpecMvPolyIntEquiv n).injective @@ -141,8 +141,8 @@ lemma hom_ext {f g : X ⟶ 𝔸(n; S)} @[reassoc] lemma comp_homOfVector {X Y : Scheme} (v : n → Γ(Y, ⊤)) (f : X ⟶ Y) (g : Y ⟶ S) : - f ≫ homOfVector g v = homOfVector (f ≫ g) (f.app ⊤ ∘ v) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + f ≫ homOfVector g v = homOfVector (f ≫ g) (f.appTop ∘ v) := by + ext1 <;> simp end homOfVector @@ -155,13 +155,13 @@ instance (v : n → Γ(X, ⊤)) : (homOfVector (X ↘ S) v).IsOver S where /-- `S`-morphisms into `Spec 𝔸(n; S)` are equivalent to the choice of `n` global sections. -/ @[simps] def homOverEquiv : { f : X ⟶ 𝔸(n; S) // f.IsOver S } ≃ (n → Γ(X, ⊤)) where - toFun f i := f.1.app ⊤ (coord S i) + toFun f i := f.1.appTop (coord S i) invFun v := ⟨homOfVector (X ↘ S) v, inferInstance⟩ left_inv f := by ext : 2 · simp [f.2.1] - · rw [homOfVector_app_top_coord] - right_inv v := by ext i; simp [-TopologicalSpace.Opens.map_top, homOfVector_app_top_coord] + · rw [homOfVector_appTop_coord] + right_inv v := by ext i; simp [-TopologicalSpace.Opens.map_top, homOfVector_appTop_coord] variable (n) in /-- @@ -171,7 +171,7 @@ Also see `AffineSpace.SpecIso`. @[simps (config := .lemmasOnly) hom inv] def isoOfIsAffine [IsAffine S] : 𝔸(n; S) ≅ Spec (.of (MvPolynomial n Γ(S, ⊤))) where - hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S)) + hom := 𝔸(n; S).toSpecΓ ≫ Spec.map (eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S)) inv := homOfVector (Spec.map C ≫ S.isoSpec.inv) ((Scheme.ΓSpecIso (.of (MvPolynomial n Γ(S, ⊤)))).inv ∘ MvPolynomial.X) hom_inv_id := by @@ -180,48 +180,43 @@ def isoOfIsAffine [IsAffine S] : rw [← Spec.map_comp_assoc, CommRingCat.comp_eq_ring_hom_comp, eval₂Hom_comp_C, ← Scheme.toSpecΓ_naturality_assoc] simp [Scheme.isoSpec] - · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, - CommRingCat.coe_comp_of, RingHom.coe_comp, Function.comp_apply, Scheme.id.base, - Scheme.id_app] - -- note: `rw [homOfVector_app_top_coord]` works but takes 3 more seconds - conv => enter [1, 2, 2]; rw [homOfVector_app_top_coord] - simp only [TopologicalSpace.Opens.map_top, Scheme.toSpecΓ_app_top, Function.comp_apply, + · simp only [Category.assoc, Scheme.comp_app, Scheme.comp_coeBase, + TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, + Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, CommRingCat.comp_apply, + homOfVector_appTop_coord, Function.comp_apply, CommRingCat.coe_of, Scheme.id_app, CommRingCat.id_apply] - erw [elementwise_of% Scheme.ΓSpecIso_naturality, Iso.hom_inv_id_apply] + erw [Iso.hom_inv_id_apply] exact eval₂_X _ _ _ inv_hom_id := by apply ext_of_isAffine simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, - TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_app_top, + TopologicalSpace.Opens.map_top, Scheme.comp_app, Scheme.toSpecΓ_appTop, Scheme.ΓSpecIso_naturality, Category.assoc, Scheme.id_app, ← Iso.eq_inv_comp, Category.comp_id] apply ringHom_ext' · show _ = CommRingCat.ofHom C ≫ _ rw [CommRingCat.comp_eq_ring_hom_comp, RingHom.comp_assoc, eval₂Hom_comp_C, ← CommRingCat.comp_eq_ring_hom_comp, ← cancel_mono (Scheme.ΓSpecIso _).hom] - erw [← Scheme.comp_app] - rw [homOfVector_over, Scheme.comp_app] - simp only [TopologicalSpace.Opens.map_top, Category.assoc, Scheme.ΓSpecIso_naturality, ← - Scheme.toSpecΓ_app_top] - erw [← Scheme.comp_app_assoc] - rw [Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] + rw [← Scheme.comp_appTop, homOfVector_over, Scheme.comp_appTop] + simp only [Category.assoc, Scheme.ΓSpecIso_naturality, ← Scheme.toSpecΓ_appTop] + rw [← Scheme.comp_appTop_assoc, Scheme.isoSpec, asIso_inv, IsIso.hom_inv_id] simp rfl · intro i erw [CommRingCat.comp_apply, coe_eval₂Hom] simp only [eval₂_X] - exact homOfVector_app_top_coord _ _ _ + exact homOfVector_appTop_coord _ _ _ @[simp] -lemma isoOfIsAffine_hom_app_top [IsAffine S] : - (isoOfIsAffine n S).hom.app ⊤ = - (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).app ⊤) (coord S) := by +lemma isoOfIsAffine_hom_appTop [IsAffine S] : + (isoOfIsAffine n S).hom.appTop = + (Scheme.ΓSpecIso _).hom ≫ eval₂Hom ((𝔸(n; S) ↘ S).appTop) (coord S) := by simp [isoOfIsAffine_hom] -@[local simp] -lemma isoOfIsAffine_inv_app_top_coord [IsAffine S] (i) : - (isoOfIsAffine n S).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma isoOfIsAffine_inv_appTop_coord [IsAffine S] (i) : + (isoOfIsAffine n S).inv.appTop (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := + homOfVector_appTop_coord _ _ _ @[reassoc (attr := simp)] lemma isoOfIsAffine_inv_over [IsAffine S] : @@ -238,27 +233,27 @@ def SpecIso (R : CommRingCat.{max u v}) : (Scheme.ΓSpecIso R).symm.commRingCatIsoToRingEquiv).toCommRingCatIso.op @[simp] -lemma SpecIso_hom_app_top (R : CommRingCat.{max u v}) : - (SpecIso n R).hom.app ⊤ = (Scheme.ΓSpecIso _).hom ≫ - eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).app ⊤) (coord (Spec R)) := by +lemma SpecIso_hom_appTop (R : CommRingCat.{max u v}) : + (SpecIso n R).hom.appTop = (Scheme.ΓSpecIso _).hom ≫ + eval₂Hom ((Scheme.ΓSpecIso _).inv ≫ (𝔸(n; Spec R) ↘ Spec R).appTop) (coord (Spec R)) := by simp only [SpecIso, Iso.trans_hom, Functor.mapIso_hom, Iso.op_hom, RingEquiv.toCommRingCatIso_hom, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, - isoOfIsAffine_hom_app_top] + isoOfIsAffine_hom_appTop] erw [Scheme.ΓSpecIso_naturality_assoc] congr 1 apply ringHom_ext' - · ext; simp; rfl + · ext; simp · simp -@[local simp] -lemma SpecIso_inv_app_top_coord (R : CommRingCat.{max u v}) (i) : - (SpecIso n R).inv.app ⊤ (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by +@[simp] +lemma SpecIso_inv_appTop_coord (R : CommRingCat.{max u v}) (i) : + (SpecIso n R).inv.appTop (coord _ i) = (Scheme.ΓSpecIso (.of _)).inv (.X i) := by simp only [SpecIso, Iso.trans_inv, Functor.mapIso_inv, Iso.op_inv, RingEquiv.toCommRingCatIso_inv, mapEquiv_symm, RingEquiv.toRingHom_eq_coe, Scheme.Spec_map, Quiver.Hom.unop_op, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] - erw [isoOfIsAffine_inv_app_top_coord, ← CommRingCat.comp_apply] + erw [isoOfIsAffine_inv_appTop_coord, ← CommRingCat.comp_apply] rw [← Scheme.ΓSpecIso_inv_naturality] erw [CommRingCat.comp_apply] congr 1 @@ -286,14 +281,14 @@ def map {S T : Scheme.{max u v}} (f : S ⟶ T) : 𝔸(n; S) ⟶ 𝔸(n; T) := lemma map_over {S T : Scheme.{max u v}} (f : S ⟶ T) : map n f ≫ 𝔸(n; T) ↘ T = 𝔸(n; S) ↘ S ≫ f := pullback.lift_fst _ _ _ -@[local simp] -lemma map_app_top_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : - (map n f).app ⊤ (coord T i) = coord S i := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma map_appTop_coord {S T : Scheme.{max u v}} (f : S ⟶ T) (i) : + (map n f).appTop (coord T i) = coord S i := + homOfVector_appTop_coord _ _ _ @[simp] lemma map_id : map n (𝟙 S) = 𝟙 𝔸(n; S) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + ext1 <;> simp @[reassoc, simp] lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : @@ -302,7 +297,7 @@ lemma map_comp {S S' S'' : Scheme} (f : S ⟶ S') (g : S' ⟶ S'') : · simp · simp only [TopologicalSpace.Opens.map_top, Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, CommRingCat.comp_apply] - erw [map_app_top_coord, map_app_top_coord, map_app_top_coord] + erw [map_appTop_coord, map_appTop_coord, map_appTop_coord] lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : map n (Spec.map φ) = @@ -314,9 +309,9 @@ lemma map_Spec_map {R S : CommRingCat.{max u v}} (φ : R ⟶ S) : rw [map_comp_C] · simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, TopologicalSpace.Opens.map_top, Scheme.comp_app, CommRingCat.comp_apply] - conv_lhs => enter[2]; tactic => exact map_app_top_coord _ _ - conv_rhs => enter[2]; tactic => exact SpecIso_inv_app_top_coord _ _ - erw [SpecIso_inv_app_top_coord, ← CommRingCat.comp_apply] + conv_lhs => enter[2]; tactic => exact map_appTop_coord _ _ + conv_rhs => enter[2]; tactic => exact SpecIso_inv_appTop_coord _ _ + erw [SpecIso_inv_appTop_coord, ← CommRingCat.comp_apply] rw [← Scheme.ΓSpecIso_inv_naturality] erw [CommRingCat.comp_apply, map_X] rfl @@ -337,39 +332,29 @@ lemma reindex_over {n m : Type v} (i : m → n) (S : Scheme.{max u v}) : reindex i S ≫ 𝔸(m; S) ↘ S = 𝔸(n; S) ↘ S := pullback.lift_fst _ _ _ -@[local simp] -lemma reindex_app_top_coord {n m : Type v} (i : m → n) (S : Scheme.{max u v}) (j : m) : - (reindex i S).app ⊤ (coord S j) = coord S (i j) := - homOfVector_app_top_coord _ _ _ +@[simp] +lemma reindex_appTop_coord {n m : Type v} (i : m → n) (S : Scheme.{max u v}) (j : m) : + (reindex i S).appTop (coord S j) = coord S (i j) := + homOfVector_appTop_coord _ _ _ @[simp] lemma reindex_id : reindex id S = 𝟙 𝔸(n; S) := by - ext1 <;> simp [-TopologicalSpace.Opens.map_top]; rfl + ext1 <;> simp +@[simp, reassoc] lemma reindex_comp {n₁ n₂ n₃ : Type v} (i : n₁ → n₂) (j : n₂ → n₃) (S : Scheme.{max u v}) : reindex (j ∘ i) S = reindex j S ≫ reindex i S := by have H₁ : reindex (j ∘ i) S ≫ 𝔸(n₁; S) ↘ S = (reindex j S ≫ reindex i S) ≫ 𝔸(n₁; S) ↘ S := by simp - have H₂ (k) : (reindex (j ∘ i) S).app ⊤ (coord S k) = - (reindex j S).app ⊤ ((reindex i S).app ⊤ (coord S k)) := by - rw [reindex_app_top_coord, reindex_app_top_coord, reindex_app_top_coord] + have H₂ (k) : (reindex (j ∘ i) S).appTop (coord S k) = + (reindex j S).appTop ((reindex i S).appTop (coord S k)) := by + rw [reindex_appTop_coord, reindex_appTop_coord, reindex_appTop_coord] rfl exact hom_ext H₁ H₂ --- These time out if added to `reindex_comp` directly. -attribute [reassoc] reindex_comp -attribute [simp] reindex_comp - +@[reassoc (attr := simp)] lemma map_reindex {n₁ n₂ : Type v} (i : n₁ → n₂) {S T : Scheme.{max u v}} (f : S ⟶ T) : map n₂ f ≫ reindex i T = reindex i S ≫ map n₁ f := by - apply hom_ext - · simp - · intro j - simp only [Scheme.comp_coeBase, TopologicalSpace.Opens.map_comp_obj, Scheme.comp_app, - CommRingCat.comp_apply, map_app_top_coord, reindex_app_top_coord] - simp only [TopologicalSpace.Opens.map_top] - erw [map_app_top_coord f (i j), reindex_app_top_coord i S] - -attribute [reassoc (attr := simp)] map_reindex + apply hom_ext <;> simp /-- The affine space as a functor. -/ @[simps] diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 6c8faa057ea86..87195c4877f61 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -415,12 +415,12 @@ theorem Scheme.toSpecΓ_base (X : Scheme.{u}) (x) : @[reassoc (attr := simp)] theorem Scheme.toSpecΓ_naturality {X Y : Scheme.{u}} (f : X ⟶ Y) : - f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.app ⊤) := + f ≫ Y.toSpecΓ = X.toSpecΓ ≫ Spec.map (f.appTop) := ΓSpec.adjunction.unit.naturality f @[simp] -theorem Scheme.toSpecΓ_app_top (X : Scheme.{u}) : - X.toSpecΓ.app ⊤ = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by +theorem Scheme.toSpecΓ_appTop (X : Scheme.{u}) : + X.toSpecΓ.appTop = (Scheme.ΓSpecIso Γ(X, ⊤)).hom := by have := ΓSpec.adjunction.left_triangle_components X dsimp at this rw [← IsIso.eq_comp_inv] at this @@ -428,6 +428,8 @@ theorem Scheme.toSpecΓ_app_top (X : Scheme.{u}) : Scheme.Γ_obj, Category.id_comp] at this rw [← Quiver.Hom.op_inj.eq_iff, this, ← op_inv, IsIso.Iso.inv_inv] +@[deprecated (since := "2024-11-23")] alias Scheme.toSpecΓ_app_top := Scheme.toSpecΓ_appTop + @[simp] theorem SpecMap_ΓSpecIso_hom (R : CommRingCat.{u}) : Spec.map ((Scheme.ΓSpecIso R).hom) = (Spec R).toSpecΓ := by @@ -437,9 +439,9 @@ theorem SpecMap_ΓSpecIso_hom (R : CommRingCat.{u}) : lemma Scheme.toSpecΓ_preimage_basicOpen (X : Scheme.{u}) (r : Γ(X, ⊤)) : X.toSpecΓ ⁻¹ᵁ (PrimeSpectrum.basicOpen r) = X.basicOpen r := by - rw [← basicOpen_eq_of_affine, Scheme.preimage_basicOpen] + rw [← basicOpen_eq_of_affine, Scheme.preimage_basicOpen, ← Scheme.Hom.appTop] congr - rw [Scheme.toSpecΓ_app_top] + rw [Scheme.toSpecΓ_appTop] exact Iso.inv_hom_id_apply _ _ -- Warning: this LHS of this lemma breaks the structure-sheaf abstraction. @@ -456,25 +458,24 @@ theorem toOpen_toSpecΓ_app {X : Scheme.{u}} (U) : exact Category.id_comp _ lemma ΓSpecIso_inv_ΓSpec_adjunction_homEquiv {X : Scheme.{u}} {B : CommRingCat} (φ : B ⟶ Γ(X, ⊤)) : - (Scheme.ΓSpecIso B).inv ≫ ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).app ⊤ = φ := by + (Scheme.ΓSpecIso B).inv ≫ ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).appTop = φ := by simp only [Adjunction.homEquiv_apply, Scheme.Spec_map, Opens.map_top, Scheme.comp_app] simp lemma ΓSpec_adjunction_homEquiv_eq {X : Scheme.{u}} {B : CommRingCat} (φ : B ⟶ Γ(X, ⊤)) : - (((ΓSpec.adjunction.homEquiv X (op B)) φ.op).app ⊤) = (Scheme.ΓSpecIso B).hom ≫ φ := by - simp_rw [← ΓSpecIso_inv_ΓSpec_adjunction_homEquiv φ] - simp + ((ΓSpec.adjunction.homEquiv X (op B)) φ.op).appTop = (Scheme.ΓSpecIso B).hom ≫ φ := by + rw [← Iso.inv_comp_eq, ΓSpecIso_inv_ΓSpec_adjunction_homEquiv] theorem ΓSpecIso_obj_hom {X : Scheme.{u}} (U : X.Opens) : - (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).app ⊤ ≫ - U.toScheme.toSpecΓ.app ⊤ ≫ U.topIso.hom := by simp + (Scheme.ΓSpecIso Γ(X, U)).hom = (Spec.map U.topIso.inv).appTop ≫ + U.toScheme.toSpecΓ.appTop ≫ U.topIso.hom := by simp @[deprecated (since := "2024-07-24")] alias ΓSpec.adjunction_unit_naturality := Scheme.toSpecΓ_naturality @[deprecated (since := "2024-07-24")] alias ΓSpec.adjunction_unit_naturality_assoc := Scheme.toSpecΓ_naturality_assoc @[deprecated (since := "2024-07-24")] -alias ΓSpec.adjunction_unit_app_app_top := Scheme.toSpecΓ_app_top +alias ΓSpec.adjunction_unit_app_app_top := Scheme.toSpecΓ_appTop @[deprecated (since := "2024-07-24")] alias ΓSpec.adjunction_unit_map_basicOpen := Scheme.toSpecΓ_preimage_basicOpen diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean index 66b28ced66e1b..10aa7de982be9 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Affine.lean @@ -158,7 +158,7 @@ instance : HasAffineProperty @IsAffineHom fun X _ _ _ ↦ IsAffine X where rw [Scheme.preimage_basicOpen] exact (isAffineOpen_top X).basicOpen _ · intro X Y _ f S hS hS' - apply_fun Ideal.map (f.app ⊤) at hS + apply_fun Ideal.map (f.appTop) at hS rw [Ideal.map_span, Ideal.map_top] at hS apply isAffine_of_isAffineOpen_basicOpen _ hS have : ∀ i : S, IsAffineOpen (f⁻¹ᵁ Y.basicOpen i.1) := hS' diff --git a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean index 6eea058bd4042..632a366ca7236 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/AffineAnd.lean @@ -35,11 +35,11 @@ variable (Q : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop /-- This is the affine target morphism property where the source is affine and the induced map of rings on global sections satisfies `P`. -/ def affineAnd : AffineTargetMorphismProperty := - fun X _ f ↦ IsAffine X ∧ Q (f.app ⊤) + fun X _ f ↦ IsAffine X ∧ Q (f.appTop) @[simp] lemma affineAnd_apply {X Y : Scheme.{u}} (f : X ⟶ Y) [IsAffine Y] : - affineAnd Q f ↔ IsAffine X ∧ Q (f.app ⊤) := + affineAnd Q f ↔ IsAffine X ∧ Q (f.appTop) := Iff.rfl attribute [local simp] AffineTargetMorphismProperty.toProperty_apply @@ -64,28 +64,29 @@ lemma affineAnd_isLocal (hPi : RingHom.RespectsIso Q) (hQl : RingHom.Localizatio constructor · simp only [Scheme.preimage_basicOpen, Opens.map_top] exact (isAffineOpen_top X).basicOpen _ - · dsimp only [Opens.map_top] - rw [morphismRestrict_app, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] + · dsimp only + rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map f (isAffineOpen_top X), - hPi.cancel_right_isIso] - haveI := (isAffineOpen_top X).isLocalization_basicOpen (f.app ⊤ r) + hPi.cancel_right_isIso, ← Scheme.Hom.appTop] + dsimp only [Opens.map_top] + haveI := (isAffineOpen_top X).isLocalization_basicOpen (f.appTop r) apply hQl exact hf of_basicOpenCover {X Y _} f s hs hf := by dsimp [affineAnd] at hf haveI : IsAffine X := by - apply isAffine_of_isAffineOpen_basicOpen (f.app ⊤ '' s) - · apply_fun Ideal.map (f.app ⊤) at hs + apply isAffine_of_isAffineOpen_basicOpen (f.appTop '' s) + · apply_fun Ideal.map (f.appTop) at hs rwa [Ideal.map_span, Ideal.map_top] at hs · rintro - ⟨r, hr, rfl⟩ simp_rw [Scheme.preimage_basicOpen] at hf exact (hf ⟨r, hr⟩).left - refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.app ⊤) s hs fun a ↦ ?_⟩ - refine ⟨Γ(Y, Y.basicOpen a.val), Γ(X, X.basicOpen (f.app ⊤ a.val)), inferInstance, + refine ⟨inferInstance, hQs.ofIsLocalization' hPi (f.appTop) s hs fun a ↦ ?_⟩ + refine ⟨Γ(Y, Y.basicOpen a.val), Γ(X, X.basicOpen (f.appTop a.val)), inferInstance, inferInstance, inferInstance, inferInstance, inferInstance, ?_, ?_⟩ - · exact (isAffineOpen_top X).isLocalization_basicOpen (f.app ⊤ a.val) + · exact (isAffineOpen_top X).isLocalization_basicOpen (f.appTop a.val) · obtain ⟨_, hf⟩ := hf a - rw [morphismRestrict_app, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] at hf + rw [morphismRestrict_appTop, hPi.cancel_right_isIso, Scheme.Opens.ι_image_top] at hf rw [(isAffineOpen_top Y).app_basicOpen_eq_away_map _ (isAffineOpen_top X)] at hf rwa [hPi.cancel_right_isIso] at hf @@ -96,7 +97,7 @@ lemma affineAnd_isStableUnderBaseChange (hQi : RingHom.RespectsIso Q) haveI : (affineAnd Q).toProperty.RespectsIso := affineAnd_respectsIso hQi apply AffineTargetMorphismProperty.IsStableUnderBaseChange.mk intro X Y S _ _ f g ⟨hY, hg⟩ - exact ⟨inferInstance, hQb.pullback_fst_app_top _ hQi f _ hg⟩ + exact ⟨inferInstance, hQb.pullback_fst_appTop _ hQi f _ hg⟩ lemma targetAffineLocally_affineAnd_iff (hQi : RingHom.RespectsIso Q) {X Y : Scheme.{u}} (f : X ⟶ Y) : @@ -133,7 +134,7 @@ lemma targetAffineLocally_affineAnd_iff_affineLocally (hQ : RingHom.PropertyIsLo intro U have : IsAffine (f ⁻¹ᵁ U) := hf.isAffine_preimage U U.2 rw [HasRingHomProperty.iff_of_isAffine (P := affineLocally Q), - morphismRestrict_app, hQ.respectsIso.cancel_right_isIso] + morphismRestrict_appTop, hQ.respectsIso.cancel_right_isIso] apply h rw [Scheme.Opens.ι_image_top] exact U.2 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean index d72195178231c..0145a3045d5db 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/ClosedImmersion.lean @@ -117,7 +117,7 @@ instance spec_of_quotient_mk {R : CommRingCat.{u}} (I : Ideal R) : /-- Any morphism between affine schemes that is surjective on global sections is a closed immersion. -/ lemma of_surjective_of_isAffine {X Y : Scheme} [IsAffine X] [IsAffine Y] (f : X ⟶ Y) - (h : Function.Surjective (f.app ⊤)) : IsClosedImmersion f := by + (h : Function.Surjective (f.appTop)) : IsClosedImmersion f := by rw [MorphismProperty.arrow_mk_iso_iff @IsClosedImmersion (arrowIsoSpecΓOfIsAffine f)] apply spec_of_surjective exact h @@ -161,7 +161,7 @@ open IsClosedImmersion LocallyRingedSpace has a closed image and `f` induces an injection on global sections, then `f` is surjective. -/ lemma surjective_of_isClosed_range_of_injective [CompactSpace X] - (hfcl : IsClosed (Set.range f.base)) (hfinj : Function.Injective (f.app ⊤)) : + (hfcl : IsClosed (Set.range f.base)) (hfinj : Function.Injective (f.appTop)) : Function.Surjective f.base := by obtain ⟨I, hI⟩ := (Scheme.eq_zeroLocus_of_isClosed_of_isAffine Y (Set.range f.base)).mp hfcl let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover @@ -186,12 +186,12 @@ lemma surjective_of_isClosed_range_of_injective [CompactSpace X] injective if it is injective on global sections. -/ lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] (hfopen : IsOpenMap f.base) (hfinj₁ : Function.Injective f.base) - (hfinj₂ : Function.Injective (f.app ⊤)) (x : X) : + (hfinj₂ : Function.Injective (f.appTop)) (x : X) : Function.Injective (f.stalkMap x) := by - let φ : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := f.app ⊤ + let φ : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := f.appTop let 𝒰 : X.OpenCover := X.affineCover.finiteSubcover have (i : 𝒰.J) : IsAffine (𝒰.obj i) := Scheme.isAffine_affineCover X _ - let res (i : 𝒰.J) : Γ(X, ⊤) ⟶ Γ(𝒰.obj i, ⊤) := (𝒰.map i).app ⊤ + let res (i : 𝒰.J) : Γ(X, ⊤) ⟶ Γ(𝒰.obj i, ⊤) := (𝒰.map i).appTop refine stalkMap_injective_of_isAffine _ _ (fun (g : Γ(Y, ⊤)) h ↦ ?_) rw [TopCat.Presheaf.Γgerm, Scheme.stalkMap_germ_apply] at h obtain ⟨U, w, (hx : x ∈ U), hg⟩ := @@ -200,8 +200,8 @@ lemma stalkMap_injective_of_isOpenMap_of_injective [CompactSpace X] (show f.base x ∈ ⟨f.base '' U.carrier, hfopen U.carrier U.is_open'⟩ from ⟨x, by simpa⟩) let W (i : 𝒰.J) : TopologicalSpace.Opens (𝒰.obj i) := (𝒰.obj i).basicOpen ((res i) (φ s)) have hwle (i : 𝒰.J) : W i ≤ (𝒰.map i)⁻¹ᵁ U := by - show ((𝒰.obj i).basicOpen ((𝒰.map i ≫ f).app ⊤ s)) ≤ _ - rw [← Scheme.preimage_basicOpen, Scheme.comp_coeBase, Opens.map_comp_obj] + show (𝒰.obj i).basicOpen ((𝒰.map i ≫ f).appTop s) ≤ _ + rw [← Scheme.preimage_basicOpen_top, Scheme.comp_coeBase, Opens.map_comp_obj] refine Scheme.Hom.preimage_le_preimage_of_le _ (le_trans (f.preimage_le_preimage_of_le bsle) (le_of_eq ?_)) simp [Set.preimage_image_eq _ hfinj₁] @@ -229,7 +229,7 @@ namespace IsClosedImmersion /-- If `f` is a closed immersion with affine target such that the induced map on global sections is injective, `f` is an isomorphism. -/ theorem isIso_of_injective_of_isAffine [IsClosedImmersion f] - (hf : Function.Injective (f.app ⊤)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| + (hf : Function.Injective (f.appTop)) : IsIso f := (isIso_iff_stalk_iso f).mpr <| have : CompactSpace X := f.isClosedEmbedding.compactSpace have hiso : IsIso f.base := TopCat.isIso_of_bijective_of_isClosedMap _ ⟨f.isClosedEmbedding.injective, @@ -244,7 +244,7 @@ variable (f) /-- If `f` is a closed immersion with affine target, the source is affine and the induced map on global sections is surjective. -/ theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : - IsAffine X ∧ Function.Surjective (f.app ⊤) := by + IsAffine X ∧ Function.Surjective (f.appTop) := by haveI i : IsClosedImmersion f := inferInstance rw [← affineTargetImageFactorization_comp f] at i ⊢ haveI := of_surjective_of_isAffine (affineTargetImageInclusion f) @@ -254,7 +254,7 @@ theorem isAffine_surjective_of_isAffine [IsClosedImmersion f] : haveI := isIso_of_injective_of_isAffine (affineTargetImageFactorization_app_injective f) exact ⟨isAffine_of_isIso (affineTargetImageFactorization f), (ConcreteCategory.bijective_of_isIso - ((affineTargetImageFactorization f).app ⊤)).surjective.comp <| + ((affineTargetImageFactorization f).appTop)).surjective.comp <| affineTargetImageInclusion_app_surjective f⟩ end IsClosedImmersion @@ -268,7 +268,7 @@ instance IsClosedImmersion.isLocalAtTarget : IsLocalAtTarget @IsClosedImmersion /-- On morphisms with affine target, being a closed immersion is precisely having affine source and being surjective on global sections. -/ instance IsClosedImmersion.hasAffineProperty : HasAffineProperty @IsClosedImmersion - (fun X _ f ↦ IsAffine X ∧ Function.Surjective (f.app ⊤)) := by + (fun X _ f ↦ IsAffine X ∧ Function.Surjective (f.appTop)) := by convert HasAffineProperty.of_isLocalAtTarget @IsClosedImmersion refine ⟨fun ⟨h₁, h₂⟩ ↦ of_surjective_of_isAffine _ h₂, by apply isAffine_surjective_of_isAffine⟩ @@ -290,7 +290,7 @@ instance IsClosedImmersion.isStableUnderBaseChange : haveI := HasAffineProperty.isLocal_affineProperty @IsClosedImmersion apply AffineTargetMorphismProperty.IsStableUnderBaseChange.mk intro X Y S _ _ f g ⟨ha, hsurj⟩ - exact ⟨inferInstance, RingHom.surjective_isStableUnderBaseChange.pullback_fst_app_top _ + exact ⟨inferInstance, RingHom.surjective_isStableUnderBaseChange.pullback_fst_appTop _ RingHom.surjective_respectsIso f _ hsurj⟩ /-- Closed immersions are locally of finite type. -/ @@ -320,13 +320,13 @@ lemma isIso_of_isClosedImmersion_of_surjective {X Y : Scheme.{u}} (f : X ⟶ Y) · infer_instance apply IsClosedImmersion.isIso_of_injective_of_isAffine obtain ⟨hX, hf⟩ := HasAffineProperty.iff_of_isAffine.mp ‹IsClosedImmersion f› - let φ := f.app ⊤ + let φ := f.appTop suffices RingHom.ker φ ≤ nilradical _ by rwa [nilradical_eq_zero, Submodule.zero_eq_bot, le_bot_iff, ← RingHom.injective_iff_ker_eq_bot] at this refine (PrimeSpectrum.zeroLocus_eq_top_iff _).mp ?_ rw [← range_specComap_of_surjective _ _ hf, Set.top_eq_univ, Set.range_eq_univ] - have : Surjective (Spec.map (f.app ⊤)) := + have : Surjective (Spec.map (f.appTop)) := (MorphismProperty.arrow_mk_iso_iff @Surjective (arrowIsoSpecΓOfIsAffine f)).mp (inferInstanceAs (Surjective f)) exact this.1 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean index 0f054e90a059e..60115fff857f2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Finite.lean @@ -37,7 +37,7 @@ class IsFinite {X Y : Scheme} (f : X ⟶ Y) extends IsAffineHom f : Prop where namespace IsFinite -instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.app ⊤)) := by +instance : HasAffineProperty @IsFinite (fun X _ f _ ↦ IsAffine X ∧ RingHom.Finite (f.appTop)) := by show HasAffineProperty @IsFinite (affineAnd RingHom.Finite) rw [HasAffineProperty.affineAnd_iff _ RingHom.finite_respectsIso RingHom.finite_localizationPreserves RingHom.finite_ofLocalizationSpan] diff --git a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean index 06cb129fa9785..ed9fb2ee4ca4d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/IsIso.lean @@ -36,10 +36,10 @@ lemma isomorphisms_eq_stalkwise : instance : IsLocalAtTarget (isomorphisms Scheme) := isomorphisms_eq_isOpenImmersion_inf_surjective ▸ inferInstance -instance : HasAffineProperty (isomorphisms Scheme) fun X _ f _ ↦ IsAffine X ∧ IsIso (f.app ⊤) := by +instance : HasAffineProperty (isomorphisms Scheme) fun X _ f _ ↦ IsAffine X ∧ IsIso (f.appTop) := by convert HasAffineProperty.of_isLocalAtTarget (isomorphisms Scheme) with X Y f hY exact ⟨fun ⟨_, _⟩ ↦ (arrow_mk_iso_iff (isomorphisms _) (arrowIsoSpecΓOfIsAffine f)).mpr - (inferInstanceAs (IsIso (Spec.map (f.app ⊤)))), + (inferInstanceAs (IsIso (Spec.map (f.appTop)))), fun (_ : IsIso f) ↦ ⟨isAffine_of_isIso f, inferInstance⟩⟩ instance : IsLocalAtTarget (monomorphisms Scheme) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean index 6a02465a5cb1a..4306631e05eb2 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/RingHomProperties.lean @@ -39,9 +39,9 @@ For `HasRingHomProperty P Q` and `f : X ⟶ Y`, we provide these API lemmas: - `AlgebraicGeometry.HasRingHomProperty.iff_appLE`: `P f` if and only if `Q (f.appLE U V _)` for all affine `U : Opens Y` and `V : Opens X`. - `AlgebraicGeometry.HasRingHomProperty.iff_of_source_openCover`: - If `Y` is affine, `P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤)` for an affine open cover `𝒰` of `X`. + If `Y` is affine, `P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop)` for an affine open cover `𝒰` of `X`. - `AlgebraicGeometry.HasRingHomProperty.iff_of_isAffine`: - If `X` and `Y` are affine, then `P f ↔ Q (f.app ⊤)`. + If `X` and `Y` are affine, then `P f ↔ Q (f.appTop)`. - `AlgebraicGeometry.HasRingHomProperty.Spec_iff`: `P (Spec.map φ) ↔ Q φ` - `AlgebraicGeometry.HasRingHomProperty.iff_of_iSup_eq_top`: @@ -66,14 +66,14 @@ namespace RingHom variable (P : ∀ {R S : Type u} [CommRing R] [CommRing S], (R →+* S) → Prop) -theorem IsStableUnderBaseChange.pullback_fst_app_top +theorem IsStableUnderBaseChange.pullback_fst_appTop (hP : IsStableUnderBaseChange P) (hP' : RespectsIso P) {X Y S : Scheme} [IsAffine X] [IsAffine Y] [IsAffine S] (f : X ⟶ S) (g : Y ⟶ S) - (H : P (g.app ⊤)) : P ((pullback.fst f g).app ⊤) := by + (H : P g.appTop) : P (pullback.fst f g).appTop := by -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11224): change `rw` to `erw` erw [← PreservesPullback.iso_inv_fst AffineScheme.forgetToScheme (AffineScheme.ofHom f) (AffineScheme.ofHom g)] - rw [Scheme.comp_app, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] + rw [Scheme.comp_appTop, hP'.cancel_right_isIso, AffineScheme.forgetToScheme_map] have := congr_arg Quiver.Hom.unop (PreservesPullback.iso_hom_fst AffineScheme.Γ.rightOp (AffineScheme.ofHom f) (AffineScheme.ofHom g)) @@ -83,6 +83,10 @@ theorem IsStableUnderBaseChange.pullback_fst_app_top rw [← this, hP'.cancel_right_isIso, ← pushoutIsoUnopPullback_inl_hom, hP'.cancel_right_isIso] exact hP.pushout_inl _ hP' _ _ H +@[deprecated (since := "2024-11-23")] +alias IsStableUnderBaseChange.pullback_fst_app_top := +IsStableUnderBaseChange.pullback_fst_appTop + end RingHom namespace AlgebraicGeometry @@ -268,10 +272,12 @@ theorem appLE (H : P f) (U : Y.affineOpens) (V : X.affineOpens) (e) : Q (f.appLE rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] at H exact H _ _ _ -theorem app_top (H : P f) [IsAffine X] [IsAffine Y] : Q (f.app ⊤) := by - rw [Scheme.Hom.app_eq_appLE] +theorem appTop (H : P f) [IsAffine X] [IsAffine Y] : Q f.appTop := by + rw [Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE] exact appLE P f H ⟨_, isAffineOpen_top _⟩ ⟨_, isAffineOpen_top _⟩ _ +@[deprecated (since := "2024-11-23")] alias app_top := appTop + include Q in theorem comp_of_isOpenImmersion [IsOpenImmersion f] (H : P g) : P (f ≫ g) := by @@ -291,7 +297,7 @@ lemma iff_appLE : P f ↔ ∀ (U : Y.affineOpens) (V : X.affineOpens) (e), Q (f. rw [eq_affineLocally P, affineLocally_iff_affineOpens_le] theorem of_source_openCover [IsAffine Y] - (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).app ⊤)) : + (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] (H : ∀ i, Q ((𝒰.map i ≫ f).appTop)) : P f := by rw [HasAffineProperty.iff_of_isAffine (P := P)] intro U @@ -313,17 +319,17 @@ theorem of_source_openCover [IsAffine Y] specialize H i rw [← (isLocal_ringHomProperty P).respectsIso.cancel_right_isIso _ ((IsOpenImmersion.isoOfRangeEq (𝒰.map i) (S i).1.ι - Subtype.range_coe.symm).inv.app _), ← Scheme.comp_app, - IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_app, - Scheme.Opens.ι_app, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H + Subtype.range_coe.symm).inv.app _), ← Scheme.comp_appTop, + IsOpenImmersion.isoOfRangeEq_inv_fac_assoc, Scheme.comp_appTop, + Scheme.Opens.ι_appTop, Scheme.Hom.appTop, Scheme.Hom.app_eq_appLE, Scheme.Hom.appLE_map] at H exact (f.appLE_congr _ rfl (by simp) Q).mp H theorem iff_of_source_openCover [IsAffine Y] (𝒰 : X.OpenCover) [∀ i, IsAffine (𝒰.obj i)] : - P f ↔ ∀ i, Q ((𝒰.map i ≫ f).app ⊤) := - ⟨fun H i ↦ app_top P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ + P f ↔ ∀ i, Q ((𝒰.map i ≫ f).appTop) := + ⟨fun H i ↦ appTop P _ (comp_of_isOpenImmersion P (𝒰.map i) f H), of_source_openCover 𝒰⟩ theorem iff_of_isAffine [IsAffine X] [IsAffine Y] : - P f ↔ Q (f.app ⊤) := by + P f ↔ Q (f.appTop) := by rw [iff_of_source_openCover (P := P) (Scheme.coverOfIsIso.{u} (𝟙 _))] simp @@ -359,7 +365,7 @@ lemma containsIdentities (hP : RingHom.ContainsIdentities Q) : P.ContainsIdentit rw [IsLocalAtTarget.iff_of_iSup_eq_top (P := P) _ (iSup_affineOpens_eq_top _)] intro U have : IsAffine (𝟙 X ⁻¹ᵁ U.1) := U.2 - rw [morphismRestrict_id, iff_of_isAffine (P := P), Scheme.id_app] + rw [morphismRestrict_id, iff_of_isAffine (P := P), Scheme.id_appTop] apply hP variable (P) in @@ -504,7 +510,7 @@ lemma isStableUnderBaseChange (hP : RingHom.IsStableUnderBaseChange Q) : limit.lift_π, PullbackCone.mk_pt, PullbackCone.mk_π_app, Category.comp_id] apply this _ (comp_of_isOpenImmersion _ _ _ H) inferInstance rw [iff_of_isAffine (P := P)] at H ⊢ - exact hP.pullback_fst_app_top _ (isLocal_ringHomProperty P).respectsIso _ _ H + exact hP.pullback_fst_appTop _ (isLocal_ringHomProperty P).respectsIso _ _ H include Q in private lemma respects_isOpenImmersion_aux diff --git a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean index 548f463af6395..66aff1bb5a74d 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/SurjectiveOnStalks.lean @@ -74,7 +74,6 @@ variable {f} in lemma iff_of_isAffine [IsAffine X] [IsAffine Y] : SurjectiveOnStalks f ↔ RingHom.SurjectiveOnStalks (f.app ⊤) := by rw [← Spec_iff, MorphismProperty.arrow_mk_iso_iff @SurjectiveOnStalks (arrowIsoSpecΓOfIsAffine f)] - rfl theorem of_comp [SurjectiveOnStalks (f ≫ g)] : SurjectiveOnStalks f := by refine ⟨fun x ↦ ?_⟩ diff --git a/Mathlib/AlgebraicGeometry/Restrict.lean b/Mathlib/AlgebraicGeometry/Restrict.lean index 4977b2dfaeffc..1d336822caba3 100644 --- a/Mathlib/AlgebraicGeometry/Restrict.lean +++ b/Mathlib/AlgebraicGeometry/Restrict.lean @@ -70,6 +70,11 @@ lemma ι_app (V) : U.ι.app V = X.presheaf.map (homOfLE (x := U.ι ''ᵁ U.ι ⁻¹ᵁ V) (Set.image_preimage_subset _ _)).op := rfl +@[simp] +lemma ι_appTop : + U.ι.appTop = X.presheaf.map (homOfLE (x := U.ι ''ᵁ ⊤) le_top).op := + rfl + @[simp] lemma ι_appLE (V W e) : U.ι.appLE V W e = @@ -170,7 +175,7 @@ def opensRestrict : instance ΓRestrictAlgebra {X : Scheme.{u}} (U : X.Opens) : Algebra (Γ(X, ⊤)) Γ(U, ⊤) := - (U.ι.app ⊤).toAlgebra + U.ι.appTop.toAlgebra lemma Scheme.map_basicOpen (r : Γ(U, ⊤)) : U.ι ''ᵁ U.toScheme.basicOpen r = X.basicOpen @@ -246,6 +251,11 @@ theorem Scheme.homOfLE_app {U V : X.Opens} (e : U ≤ V) (W : Opens V) : rw [e₃, ← Functor.map_comp] congr 1 +theorem Scheme.homOfLE_appTop {U V : X.Opens} (e : U ≤ V) : + (X.homOfLE e).appTop = + X.presheaf.map (homOfLE <| X.ι_image_homOfLE_le_ι_image e ⊤).op := + homOfLE_app .. + instance (X : Scheme.{u}) {U V : X.Opens} (e : U ≤ V) : IsOpenImmersion (X.homOfLE e) := by delta Scheme.homOfLE infer_instance @@ -295,7 +305,7 @@ def Scheme.restrictFunctorΓ : X.restrictFunctor.op ⋙ (Over.forget X).op ⋙ S (by intro U V i dsimp - rw [X.homOfLE_app, ← Functor.map_comp, ← Functor.map_comp] + rw [X.homOfLE_appTop, ← Functor.map_comp, ← Functor.map_comp] congr 1) /-- `X ∣_ U ∣_ V` is isomorphic to `X ∣_ V ∣_ U` -/ @@ -531,6 +541,11 @@ theorem morphismRestrict_app {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Quiver.Hom.unop_op, Hom.opensFunctor_map_homOfLE] rw [this, Hom.appLE_map, Hom.appLE_map, Hom.appLE_map] +theorem morphismRestrict_appTop {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : + (f ∣_ U).appTop = f.app (U.ι ''ᵁ ⊤) ≫ + X.presheaf.map (eqToHom (image_morphismRestrict_preimage f U ⊤)).op := + morphismRestrict_app .. + @[simp] theorem morphismRestrict_app' {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) (V : Opens U) : (f ∣_ U).app V = f.appLE _ _ (image_morphismRestrict_preimage f U V).le := @@ -547,7 +562,7 @@ theorem Γ_map_morphismRestrict {X Y : Scheme.{u}} (f : X ⟶ Y) (U : Y.Opens) : Scheme.Γ.map (f ∣_ U).op = Y.presheaf.map (eqToHom U.isOpenEmbedding_obj_top.symm).op ≫ f.app U ≫ X.presheaf.map (eqToHom (f ⁻¹ᵁ U).isOpenEmbedding_obj_top).op := by - rw [Scheme.Γ_map_op, morphismRestrict_app f U ⊤, f.naturality_assoc, ← X.presheaf.map_comp] + rw [Scheme.Γ_map_op, morphismRestrict_appTop f U, f.naturality_assoc, ← X.presheaf.map_comp] rfl /-- Restricting a morphism onto the image of an open immersion is isomorphic to the base change @@ -694,11 +709,11 @@ end Scheme.Hom /-- `f.resLE U V` induces `f.appLE U V` on global sections. -/ noncomputable def arrowResLEAppIso (f : X ⟶ Y) (U : Y.Opens) (V : X.Opens) (e : V ≤ f ⁻¹ᵁ U) : - Arrow.mk ((f.resLE U V e).app ⊤) ≅ Arrow.mk (f.appLE U V e) := + Arrow.mk ((f.resLE U V e).appTop) ≅ Arrow.mk (f.appLE U V e) := Arrow.isoMk U.topIso V.topIso <| by simp only [Opens.map_top, Arrow.mk_left, Arrow.mk_right, Functor.id_obj, Scheme.Opens.topIso_hom, eqToHom_op, Arrow.mk_hom, Scheme.Hom.map_appLE] - rw [← Scheme.Hom.appLE_eq_app, Scheme.Hom.resLE_appLE, Scheme.Hom.appLE_map] + rw [Scheme.Hom.appTop, ← Scheme.Hom.appLE_eq_app, Scheme.Hom.resLE_appLE, Scheme.Hom.appLE_map] end MorphismRestrict diff --git a/Mathlib/AlgebraicGeometry/Scheme.lean b/Mathlib/AlgebraicGeometry/Scheme.lean index 6decb4e3cb660..dc282d80303ad 100644 --- a/Mathlib/AlgebraicGeometry/Scheme.lean +++ b/Mathlib/AlgebraicGeometry/Scheme.lean @@ -108,6 +108,11 @@ this is the induced map `Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U)`. -/ abbrev app (U : Y.Opens) : Γ(Y, U) ⟶ Γ(X, f ⁻¹ᵁ U) := f.c.app (op U) +/-- Given a morphism of schemes `f : X ⟶ Y`, +this is the induced map `Γ(Y, ⊤) ⟶ Γ(X, ⊤)`. -/ +abbrev appTop : Γ(Y, ⊤) ⟶ Γ(X, ⊤) := + f.app ⊤ + @[reassoc] lemma naturality (i : op U' ⟶ op U) : Y.presheaf.map i ≫ f.app U = f.app U' ≫ X.presheaf.map ((Opens.map f.base).map i.unop).op := @@ -250,6 +255,11 @@ theorem id.base (X : Scheme) : (𝟙 X : _).base = 𝟙 _ := theorem id_app {X : Scheme} (U : X.Opens) : (𝟙 X : _).app U = 𝟙 _ := rfl +@[simp] +theorem id_appTop {X : Scheme} : + (𝟙 X : _).appTop = 𝟙 _ := + rfl + @[reassoc] theorem comp_toLRSHom {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).toLRSHom = f.toLRSHom ≫ g.toLRSHom := @@ -275,6 +285,11 @@ theorem comp_app {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) (U) : (f ≫ g).app U = g.app U ≫ f.app _ := rfl +@[simp, reassoc] -- reassoc lemma does not need `simp` +theorem comp_appTop {X Y Z : Scheme} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).appTop = g.appTop ≫ f.appTop := + rfl + @[deprecated (since := "2024-06-23")] alias comp_val_c_app := comp_app @[deprecated (since := "2024-06-23")] alias comp_val_c_app_assoc := comp_app_assoc @@ -335,8 +350,10 @@ theorem inv_app {X Y : Scheme} (f : X ⟶ Y) [IsIso f] (U : X.Opens) : rw [IsIso.eq_comp_inv, ← Scheme.comp_app, Scheme.congr_app (IsIso.hom_inv_id f), Scheme.id_app, Category.id_comp] -theorem inv_app_top {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : - (inv f).app ⊤ = inv (f.app ⊤) := by simp +theorem inv_appTop {X Y : Scheme} (f : X ⟶ Y) [IsIso f] : + (inv f).appTop = inv (f.appTop) := by simp + +@[deprecated (since := "2024-11-23")] alias inv_app_top := inv_appTop end Scheme @@ -440,10 +457,10 @@ theorem Γ_obj_op (X : Scheme) : Γ.obj (op X) = Γ(X, ⊤) := rfl @[simp] -theorem Γ_map {X Y : Schemeᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.app ⊤ := +theorem Γ_map {X Y : Schemeᵒᵖ} (f : X ⟶ Y) : Γ.map f = f.unop.appTop := rfl -theorem Γ_map_op {X Y : Scheme} (f : X ⟶ Y) : Γ.map f.op = f.app ⊤ := +theorem Γ_map_op {X Y : Scheme} (f : X ⟶ Y) : Γ.map f.op = f.appTop := rfl /-- @@ -466,13 +483,13 @@ def ΓSpecIso : Γ(Spec R, ⊤) ≅ R := SpecΓIdentity.app R @[reassoc (attr := simp)] lemma ΓSpecIso_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : - (Spec.map f).app ⊤ ≫ (ΓSpecIso S).hom = (ΓSpecIso R).hom ≫ f := SpecΓIdentity.hom.naturality f + (Spec.map f).appTop ≫ (ΓSpecIso S).hom = (ΓSpecIso R).hom ≫ f := SpecΓIdentity.hom.naturality f -- The RHS is not necessarily simpler than the LHS, but this direction coincides with the simp -- direction of `NatTrans.naturality`. @[reassoc (attr := simp)] lemma ΓSpecIso_inv_naturality {R S : CommRingCat.{u}} (f : R ⟶ S) : - f ≫ (ΓSpecIso S).inv = (ΓSpecIso R).inv ≫ (Spec.map f).app ⊤ := SpecΓIdentity.inv.naturality f + f ≫ (ΓSpecIso S).inv = (ΓSpecIso R).inv ≫ (Spec.map f).appTop := SpecΓIdentity.inv.naturality f -- This is not marked simp to respect the abstraction lemma ΓSpecIso_inv : (ΓSpecIso R).inv = StructureSheaf.toOpen R ⊤ := rfl @@ -538,6 +555,10 @@ theorem preimage_basicOpen {X Y : Scheme.{u}} (f : X ⟶ Y) {U : Y.Opens} (r : f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.app U r) := LocallyRingedSpace.preimage_basicOpen f.toLRSHom r +theorem preimage_basicOpen_top {X Y : Scheme.{u}} (f : X ⟶ Y) (r : Γ(Y, ⊤)) : + f ⁻¹ᵁ (Y.basicOpen r) = X.basicOpen (f.appTop r) := + preimage_basicOpen .. + lemma basicOpen_appLE {X Y : Scheme.{u}} (f : X ⟶ Y) (U : X.Opens) (V : Y.Opens) (e : U ≤ f ⁻¹ᵁ V) (s : Γ(Y, V)) : X.basicOpen (f.appLE V U e s) = U ⊓ f ⁻¹ᵁ (Y.basicOpen s) := by simp only [preimage_basicOpen, Hom.appLE, CommRingCat.coe_comp_of, RingHom.coe_comp, diff --git a/Mathlib/AlgebraicGeometry/Stalk.lean b/Mathlib/AlgebraicGeometry/Stalk.lean index 6f520c2514a8f..e06eb1b7ef0ca 100644 --- a/Mathlib/AlgebraicGeometry/Stalk.lean +++ b/Mathlib/AlgebraicGeometry/Stalk.lean @@ -112,6 +112,13 @@ lemma fromSpecStalk_app {x : X} (hxU : x ∈ U) : hV.fromSpec_app_of_le _ hVU, ← X.presheaf.germ_res (homOfLE hVU) x hxV] simp [Category.assoc, ← ΓSpecIso_inv_naturality_assoc] +lemma fromSpecStalk_appTop {x : X} : + (X.fromSpecStalk x).appTop = + X.presheaf.germ ⊤ x trivial ≫ + (ΓSpecIso (X.presheaf.stalk x)).inv ≫ + (Spec (X.presheaf.stalk x)).presheaf.map (homOfLE le_top).op := + fromSpecStalk_app .. + @[reassoc (attr := simp)] lemma Spec_map_stalkSpecializes_fromSpecStalk {x y : X} (h : x ⤳ y) : Spec.map (X.presheaf.stalkSpecializes h) ≫ X.fromSpecStalk y = X.fromSpecStalk x := by @@ -184,7 +191,7 @@ instance {X : Scheme.{u}} (U : X.Opens) (x : X) (hxU : x ∈ U) : lemma fromSpecStalk_toSpecΓ (X : Scheme.{u}) (x : X) : X.fromSpecStalk x ≫ X.toSpecΓ = Spec.map (X.presheaf.germ ⊤ x trivial) := by rw [Scheme.toSpecΓ_naturality, ← SpecMap_ΓSpecIso_hom, ← Spec.map_comp, - @Scheme.fromSpecStalk_app X ⊤ _ trivial] + Scheme.fromSpecStalk_appTop] simp @[reassoc (attr := simp)] diff --git a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean index cb9082af37462..dbc4ac5a16007 100644 --- a/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean +++ b/Mathlib/AlgebraicGeometry/ValuativeCriterion.lean @@ -275,7 +275,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique conv_lhs => rw [← pullback.lift_fst l₁ l₂ h, ← pullback.condition_assoc] conv_rhs => rw [← pullback.lift_snd l₁ l₂ h, ← pullback.condition_assoc] simp - suffices h : Function.Bijective (g.app ⊤) by + suffices h : Function.Bijective (g.appTop) by refine (HasAffineProperty.iff_of_isAffine (P := MorphismProperty.isomorphisms Scheme)).mpr ?_ exact ⟨hZ, (ConcreteCategory.isIso_iff_bijective _).mpr h⟩ constructor @@ -286,7 +286,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique simpa using hl₂.symm have hg : l ≫ g = Spec.map (CommRingCat.ofHom (algebraMap S.R S.K)) := pullback.lift_snd _ _ _ - have : Function.Injective ((l ≫ g).app ⊤) := by + have : Function.Injective ((l ≫ g).appTop) := by rw [hg] let e := arrowIsoΓSpecOfIsAffine (CommRingCat.ofHom <| algebraMap S.R S.K) let P : MorphismProperty CommRingCat := @@ -296,7 +296,7 @@ lemma IsSeparated.valuativeCriterion [IsSeparated f] : ValuativeCriterion.Unique show P _ rw [← MorphismProperty.arrow_mk_iso_iff (P := P) e] exact NoZeroSMulDivisors.algebraMap_injective S.R S.K - rw [Scheme.comp_app _ _] at this + rw [Scheme.comp_appTop] at this exact Function.Injective.of_comp this · rw [@HasAffineProperty.iff_of_isAffine @IsClosedImmersion] at this exact this.right From 835c2946b90defd880727950cbcb678d11386e9e Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sun, 24 Nov 2024 18:12:45 +0000 Subject: [PATCH 105/250] Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/6200 --- 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 b9d6d041d6764..503457238bd1d 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "922825a3fa019deafbf4f61ed7c70fbdeebc9206", + "rev": "c7de4e5f4b3ec901cbe6340134af58b2c609902c", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "nightly-testing", + "inputRev": "lean-pr-testing-6200", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index 5c55a15c5c5e2..3298fa371855b 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-6200" 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 80611d83b2915..005ca7b7ab168 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-19 +leanprover/lean4-pr-releases:pr-release-6200 From f28fa7213f2fe9065a3a12f88d256ea7dd006d7e Mon Sep 17 00:00:00 2001 From: Pieter Cuijpers Date: Sun, 24 Nov 2024 20:04:13 +0000 Subject: [PATCH 106/250] feat: introduce the missing notation for `AddHom` (#19387) I noticed that there is notation for a non-unital monoid hom but not for a non-unital additive monoid hom, and since I'm working on extensions of non-unital homomorphisms in my PR's on quantale, I found it natural to add this. --- Mathlib/Algebra/Group/Hom/Defs.lean | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Hom/Defs.lean b/Mathlib/Algebra/Group/Hom/Defs.lean index f2540ca0f00ab..ea061a0993f08 100644 --- a/Mathlib/Algebra/Group/Hom/Defs.lean +++ b/Mathlib/Algebra/Group/Hom/Defs.lean @@ -30,6 +30,7 @@ building blocks for other homomorphisms: * `→+`: Bundled `AddMonoid` homs. Also use for `AddGroup` homs. * `→*`: Bundled `Monoid` homs. Also use for `Group` homs. +* `→ₙ+`: Bundled `AddSemigroup` homs. * `→ₙ*`: Bundled `Semigroup` homs. ## Implementation notes @@ -95,7 +96,9 @@ end Zero section Add -/-- `AddHom M N` is the type of functions `M → N` that preserve addition. +/-- `M →ₙ+ N` is the type of functions `M → N` that preserve addition. The `ₙ` in the notation +stands for "non-unital" because it is intended to match the notation for `NonUnitalAlgHom` and +`NonUnitalRingHom`, so a `AddHom` is a non-unital additive monoid hom. When possible, instead of parametrizing results over `(f : AddHom M N)`, you should parametrize over `(F : Type*) [AddHomClass F M N] (f : F)`. @@ -108,6 +111,9 @@ structure AddHom (M : Type*) (N : Type*) [Add M] [Add N] where /-- The proposition that the function preserves addition -/ protected map_add' : ∀ x y, toFun (x + y) = toFun x + toFun y +/-- `M →ₙ+ N` denotes the type of addition-preserving maps from `M` to `N`. -/ +infixr:25 " →ₙ+ " => AddHom + /-- `AddHomClass F M N` states that `F` is a type of addition-preserving homomorphisms. You should declare an instance of this typeclass when you extend `AddHom`. -/ From 64cdba7b65f47ec7747a1cf6667c3a3dd5ed6e15 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Sun, 24 Nov 2024 21:16:49 +0000 Subject: [PATCH 107/250] Trigger CI for https://github.com/leanprover/lean4/pull/6200 --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 503457238bd1d..4c5c1f493c3f1 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "c7de4e5f4b3ec901cbe6340134af58b2c609902c", + "rev": "a7187480117d6141fb4c46d1c677544037e35c7d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6200", From 7dd730c61e3a593c7add11a30644ca66ed91eeee Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Mon, 25 Nov 2024 00:31:59 +0000 Subject: [PATCH 108/250] Trigger CI for https://github.com/leanprover/lean4/pull/6200 --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 4c5c1f493c3f1..785a40b7d3e31 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "a7187480117d6141fb4c46d1c677544037e35c7d", + "rev": "4ab397409a4f3273085129d922f6c27ebf299621", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6200", From 1cc7f6f91c168d38513b2946b9aee80e20fd0f54 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Mon, 25 Nov 2024 00:46:00 +0000 Subject: [PATCH 109/250] 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 044c8fb955bab..0e0184ac9c4d6 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "85cadff63aab36f19c18434921628449ab540bd6", + "rev": "dc439388cc078925934b1c3e06276194280822d8", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From 5a054eff5942f119027ee221b975c4488dda9ed3 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Mon, 25 Nov 2024 01:06:41 +0000 Subject: [PATCH 110/250] chore: update Mathlib dependencies 2024-11-25 (#19446) 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 6e18dc8970e13..aa466b92f0770 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0dc51ac7947ff6aa2c16bcffb64c46c7149d1276", + "rev": "44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From bb8b98c3afdef41dc718b24d52a149eb8050a611 Mon Sep 17 00:00:00 2001 From: tydeu Date: Sun, 24 Nov 2024 18:20:34 -0500 Subject: [PATCH 111/250] chore: adaptions for leanprover/lean4#6200 --- Mathlib/Algebra/IsPrimePow.lean | 2 +- Mathlib/Combinatorics/Colex.lean | 2 +- Mathlib/Computability/Ackermann.lean | 2 +- Mathlib/Data/Nat/BitIndices.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 15 +++------------ Mathlib/Data/Nat/Factorization/Basic.lean | 4 ++-- Mathlib/Data/Nat/Log.lean | 2 +- Mathlib/Dynamics/Newton.lean | 2 +- Mathlib/FieldTheory/IsPerfectClosure.lean | 2 +- Mathlib/FieldTheory/PerfectClosure.lean | 2 +- Mathlib/NumberTheory/FactorisationProperties.lean | 2 +- Mathlib/NumberTheory/FermatPsp.lean | 2 +- Mathlib/NumberTheory/Padics/PadicIntegers.lean | 2 +- Mathlib/RingTheory/Multiplicity.lean | 2 +- Mathlib/RingTheory/WittVector/Frobenius.lean | 2 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 2 +- 16 files changed, 19 insertions(+), 28 deletions(-) diff --git a/Mathlib/Algebra/IsPrimePow.lean b/Mathlib/Algebra/IsPrimePow.lean index c21cf0363fe86..d9de71addc9b7 100644 --- a/Mathlib/Algebra/IsPrimePow.lean +++ b/Mathlib/Algebra/IsPrimePow.lean @@ -71,7 +71,7 @@ theorem isPrimePow_nat_iff_bounded (n : ℕ) : rw [isPrimePow_nat_iff] refine Iff.symm ⟨fun ⟨p, _, k, _, hp, hk, hn⟩ => ⟨p, k, hp, hk, hn⟩, ?_⟩ rintro ⟨p, k, hp, hk, rfl⟩ - refine ⟨p, ?_, k, (Nat.lt_pow_self hp.one_lt _).le, hp, hk, rfl⟩ + refine ⟨p, ?_, k, (Nat.lt_pow_self hp.one_lt).le, hp, hk, rfl⟩ conv => { lhs; rw [← (pow_one p)] } exact Nat.pow_le_pow_right hp.one_lt.le hk diff --git a/Mathlib/Combinatorics/Colex.lean b/Mathlib/Combinatorics/Colex.lean index b828869fcbf1e..7c732db7e7473 100644 --- a/Mathlib/Combinatorics/Colex.lean +++ b/Mathlib/Combinatorics/Colex.lean @@ -534,7 +534,7 @@ theorem geomSum_injective {n : ℕ} (hn : 2 ≤ n) : geomSum_le_geomSum_iff_toColex_le_toColex hn, ← le_antisymm_iff, Colex.toColex.injEq] at h theorem lt_geomSum_of_mem {a : ℕ} (hn : 2 ≤ n) (hi : a ∈ s) : a < ∑ i in s, n ^ i := - (Nat.lt_pow_self hn a).trans_le <| single_le_sum (by simp) hi + (a.lt_pow_self hn).trans_le <| single_le_sum (by simp) hi @[simp] theorem toFinset_bitIndices_twoPowSum (s : Finset ℕ) : (∑ i in s, 2 ^ i).bitIndices.toFinset = s := by diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index ace0d47e3ed95..84d4c197c5587 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -235,7 +235,7 @@ private theorem sq_le_two_pow_add_one_minus_three (n : ℕ) : n ^ 2 ≤ 2 ^ (n + norm_num apply succ_le_of_lt rw [Nat.pow_succ, mul_comm _ 2, mul_lt_mul_left (zero_lt_two' ℕ)] - apply lt_two_pow + exact Nat.lt_two_pow · rw [Nat.pow_succ, Nat.pow_succ] linarith [one_le_pow k 2 zero_lt_two] diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 9d9c7c094edad..659563741e3e6 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -114,6 +114,6 @@ theorem two_pow_le_of_mem_bitIndices (ha : a ∈ n.bitIndices) : 2^a ≤ n := by exact List.single_le_sum (by simp) _ <| mem_map_of_mem _ ha theorem not_mem_bitIndices_self (n : ℕ) : n ∉ n.bitIndices := - fun h ↦ (lt_two_pow n).not_le <| two_pow_le_of_mem_bitIndices h + fun h ↦ (n.lt_two_pow).not_le <| two_pow_le_of_mem_bitIndices h end Nat diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 9063d971434ff..3fe52d5e22983 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -615,15 +615,6 @@ lemma le_self_pow (hn : n ≠ 0) : ∀ a : ℕ, a ≤ a ^ n | 0 => zero_le _ | a + 1 => by simpa using Nat.pow_le_pow_right a.succ_pos (Nat.one_le_iff_ne_zero.2 hn) -lemma lt_pow_self (ha : 1 < a) : ∀ n : ℕ, n < a ^ n - | 0 => by simp - | n + 1 => - calc - n + 1 < a ^ n + 1 := Nat.add_lt_add_right (lt_pow_self ha _) _ - _ ≤ a ^ (n + 1) := Nat.pow_lt_pow_succ ha - -lemma lt_two_pow (n : ℕ) : n < 2 ^ n := lt_pow_self (by decide) n - lemma one_le_pow (n m : ℕ) (h : 0 < m) : 1 ≤ m ^ n := by simpa using Nat.pow_le_pow_of_le_left h n lemma one_le_pow' (n m : ℕ) : 1 ≤ (m + 1) ^ n := one_le_pow n (m + 1) (succ_pos m) @@ -646,7 +637,7 @@ lemma one_lt_two_pow' (n : ℕ) : 1 < 2 ^ (n + 1) := one_lt_pow n.succ_ne_zero ( lemma mul_lt_mul_pow_succ (ha : 0 < a) (hb : 1 < b) : n * b < a * b ^ (n + 1) := by rw [Nat.pow_succ, ← Nat.mul_assoc, Nat.mul_lt_mul_right (Nat.lt_trans Nat.zero_lt_one hb)] exact Nat.lt_of_le_of_lt (Nat.le_mul_of_pos_left _ ha) - ((Nat.mul_lt_mul_left ha).2 <| Nat.lt_pow_self hb _) + ((Nat.mul_lt_mul_left ha).2 <| Nat.lt_pow_self hb) lemma sq_sub_sq (a b : ℕ) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by simpa [Nat.pow_succ] using Nat.mul_self_sub_mul_self_eq a b @@ -1023,7 +1014,7 @@ lemma not_pos_pow_dvd : ∀ {a n : ℕ} (_ : 1 < a) (_ : 1 < n), ¬ a ^ n ∣ a have : succ a * succ a ^ n ∣ succ a * 1 := by simpa [pow_succ'] using h have : succ a ^ n ∣ 1 := Nat.dvd_of_mul_dvd_mul_left (succ_pos _) this have he : succ a ^ n = 1 := eq_one_of_dvd_one this - have : n < succ a ^ n := lt_pow_self hp n + have : n < succ a ^ n := n.lt_pow_self hp have : n < 1 := by rwa [he] at this have : n = 0 := Nat.eq_zero_of_le_zero <| le_of_lt_succ this have : 1 < 1 := by rwa [this] at hk @@ -1031,7 +1022,7 @@ lemma not_pos_pow_dvd : ∀ {a n : ℕ} (_ : 1 < a) (_ : 1 < n), ¬ a ^ n ∣ a lemma lt_of_pow_dvd_right (hb : b ≠ 0) (ha : 2 ≤ a) (h : a ^ n ∣ b) : n < b := by rw [← Nat.pow_lt_pow_iff_right (succ_le_iff.1 ha)] - exact Nat.lt_of_le_of_lt (le_of_dvd (Nat.pos_iff_ne_zero.2 hb) h) (lt_pow_self ha _) + exact Nat.lt_of_le_of_lt (le_of_dvd (Nat.pos_iff_ne_zero.2 hb) h) (Nat.lt_pow_self ha) lemma div_dvd_of_dvd (h : n ∣ m) : m / n ∣ m := ⟨n, (Nat.div_mul_cancel h).symm⟩ diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 4e4bd0fa951c0..73b4652c0c4a3 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -155,7 +155,7 @@ theorem ordCompl_mul (a b p : ℕ) : ordCompl[p] (a * b) = ordCompl[p] a * ordCo theorem factorization_lt {n : ℕ} (p : ℕ) (hn : n ≠ 0) : n.factorization p < n := by by_cases pp : p.Prime · exact (Nat.pow_lt_pow_iff_right pp.one_lt).1 <| (ordProj_le p hn).trans_lt <| - lt_pow_self pp.one_lt _ + Nat.lt_pow_self pp.one_lt · simpa only [factorization_eq_zero_of_non_prime n pp] using hn.bot_lt /-- An upper bound on `n.factorization p` -/ @@ -356,7 +356,7 @@ theorem dvd_iff_prime_pow_dvd_dvd (n d : ℕ) : · simp rcases eq_or_ne d 0 with (rfl | hd) · simp only [zero_dvd_iff, hn, false_iff, not_forall] - exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (lt_two_pow n).not_le⟩ + exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (n.lt_two_pow).not_le⟩ refine ⟨fun h p k _ hpkd => dvd_trans hpkd h, ?_⟩ rw [← factorization_prime_le_iff_dvd hd hn] intro h p pp diff --git a/Mathlib/Data/Nat/Log.lean b/Mathlib/Data/Nat/Log.lean index fc11e395a0c41..90661f042fa2c 100644 --- a/Mathlib/Data/Nat/Log.lean +++ b/Mathlib/Data/Nat/Log.lean @@ -123,7 +123,7 @@ theorem lt_pow_of_log_lt {b x y : ℕ} (hb : 1 < b) : log b y < x → y < b ^ x lemma log_lt_self (b : ℕ) {x : ℕ} (hx : x ≠ 0) : log b x < x := match le_or_lt b 1 with | .inl h => log_of_left_le_one h x ▸ Nat.pos_iff_ne_zero.2 hx - | .inr h => log_lt_of_lt_pow hx <| lt_pow_self h _ + | .inr h => log_lt_of_lt_pow hx <| Nat.lt_pow_self h lemma log_le_self (b x : ℕ) : log b x ≤ x := if hx : x = 0 then by simp [hx] diff --git a/Mathlib/Dynamics/Newton.lean b/Mathlib/Dynamics/Newton.lean index 94030f064065f..3d992313b0958 100644 --- a/Mathlib/Dynamics/Newton.lean +++ b/Mathlib/Dynamics/Newton.lean @@ -111,7 +111,7 @@ theorem exists_unique_nilpotent_sub_and_aeval_eq_zero · -- Existence obtain ⟨n, hn⟩ := id h refine ⟨P.newtonMap^[n] x, isNilpotent_iterate_newtonMap_sub_of_isNilpotent h n, ?_⟩ - rw [← zero_dvd_iff, ← pow_eq_zero_of_le n.lt_two_pow.le hn] + rw [← zero_dvd_iff, ← pow_eq_zero_of_le (n.lt_two_pow).le hn] exact aeval_pow_two_pow_dvd_aeval_iterate_newtonMap h h' n · -- Uniqueness have ⟨u, hu⟩ := binomExpansion (P.map (algebraMap R S)) r₁ (r₂ - r₁) diff --git a/Mathlib/FieldTheory/IsPerfectClosure.lean b/Mathlib/FieldTheory/IsPerfectClosure.lean index 5c6d7928dc1a4..db75df108ee8c 100644 --- a/Mathlib/FieldTheory/IsPerfectClosure.lean +++ b/Mathlib/FieldTheory/IsPerfectClosure.lean @@ -96,7 +96,7 @@ theorem mem_pNilradical {R : Type*} [CommSemiring R] {p : ℕ} {x : R} : by_cases hp : 1 < p · rw [pNilradical_eq_nilradical hp] refine ⟨fun ⟨n, h⟩ ↦ ⟨n, ?_⟩, fun ⟨n, h⟩ ↦ ⟨p ^ n, h⟩⟩ - rw [← Nat.sub_add_cancel ((Nat.lt_pow_self hp n).le), pow_add, h, mul_zero] + rw [← Nat.sub_add_cancel ((n.lt_pow_self hp).le), pow_add, h, mul_zero] rw [pNilradical_eq_bot hp, Ideal.mem_bot] refine ⟨fun h ↦ ⟨0, by rw [pow_zero, pow_one, h]⟩, fun ⟨n, h⟩ ↦ ?_⟩ rcases Nat.le_one_iff_eq_zero_or_eq_one.1 (not_lt.1 hp) with hp | hp diff --git a/Mathlib/FieldTheory/PerfectClosure.lean b/Mathlib/FieldTheory/PerfectClosure.lean index 6b6f45e166246..7f553815a9734 100644 --- a/Mathlib/FieldTheory/PerfectClosure.lean +++ b/Mathlib/FieldTheory/PerfectClosure.lean @@ -404,7 +404,7 @@ theorem of_apply (x : K) : of K p x = mk _ _ (0, x) := instance instReduced : IsReduced (PerfectClosure K p) where eq_zero x := induction_on x fun x ⟨n, h⟩ ↦ by replace h : mk K p x ^ p ^ n = 0 := by - rw [← Nat.sub_add_cancel ((Nat.lt_pow_self (Fact.out : p.Prime).one_lt n).le), + rw [← Nat.sub_add_cancel ((n.lt_pow_self (Fact.out : p.Prime).one_lt).le), pow_add, h, mul_zero] simp only [zero_def, mk_pow, mk_eq_iff, zero_add, ← coe_iterateFrobenius, map_zero] at h ⊢ obtain ⟨m, h⟩ := h diff --git a/Mathlib/NumberTheory/FactorisationProperties.lean b/Mathlib/NumberTheory/FactorisationProperties.lean index 28cabcc7fd958..206eadd98f223 100644 --- a/Mathlib/NumberTheory/FactorisationProperties.lean +++ b/Mathlib/NumberTheory/FactorisationProperties.lean @@ -185,7 +185,7 @@ theorem infinite_even_deficient : {n : ℕ | Even n ∧ n.Deficient}.Infinite := constructor · exact ⟨⟨2 ^ n, by ring⟩, prime_two.deficient_pow⟩ · calc - n ≤ 2 ^ n := Nat.le_of_lt (lt_two_pow n) + n ≤ 2 ^ n := Nat.le_of_lt n.lt_two_pow _ < 2 ^ (n + 1) := (Nat.pow_lt_pow_iff_right (Nat.one_lt_two)).mpr (lt_add_one n) theorem infinite_odd_deficient : {n : ℕ | Odd n ∧ n.Deficient}.Infinite := by diff --git a/Mathlib/NumberTheory/FermatPsp.lean b/Mathlib/NumberTheory/FermatPsp.lean index 9fc7ee0822135..0e9039ba0b808 100644 --- a/Mathlib/NumberTheory/FermatPsp.lean +++ b/Mathlib/NumberTheory/FermatPsp.lean @@ -316,7 +316,7 @@ private theorem psp_from_prime_gt_p {b : ℕ} (b_ge_two : 2 ≤ b) {p : ℕ} (p_ have : 2 ≤ 2 * p - 2 := le_tsub_of_add_le_left (show 4 ≤ 2 * p by omega) have : 2 + p ≤ 2 * p := by omega have : p ≤ 2 * p - 2 := le_tsub_of_add_le_left this - exact this.trans_lt (lt_pow_self b_ge_two _) + exact this.trans_lt (Nat.lt_pow_self b_ge_two) /-- For all positive bases, there exist infinite **Fermat pseudoprimes** to that base. Given in this form: for all numbers `b ≥ 1` and `m`, there exists a pseudoprime `n` to base `b` such diff --git a/Mathlib/NumberTheory/Padics/PadicIntegers.lean b/Mathlib/NumberTheory/Padics/PadicIntegers.lean index 53e32586436a1..a376d9382a73e 100644 --- a/Mathlib/NumberTheory/Padics/PadicIntegers.lean +++ b/Mathlib/NumberTheory/Padics/PadicIntegers.lean @@ -309,7 +309,7 @@ theorem exists_pow_neg_lt {ε : ℝ} (hε : 0 < ε) : ∃ k : ℕ, (p : ℝ) ^ ( apply lt_of_lt_of_le hk norm_cast apply le_of_lt - convert Nat.lt_pow_self _ _ using 1 + convert Nat.lt_pow_self _ using 1 exact hp.1.one_lt · exact mod_cast hp.1.pos diff --git a/Mathlib/RingTheory/Multiplicity.lean b/Mathlib/RingTheory/Multiplicity.lean index 85ab17ad0be89..80b16bb391718 100644 --- a/Mathlib/RingTheory/Multiplicity.lean +++ b/Mathlib/RingTheory/Multiplicity.lean @@ -413,7 +413,7 @@ theorem Nat.multiplicity_finite_iff {a b : ℕ} : Finite a b ↔ a ≠ 1 ∧ 0 < | 0 => ha rfl | 1 => ha1 rfl | b+2 => by omega - not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (lt_pow_self ha_gt_one b), + not_lt_of_ge (le_of_dvd (Nat.pos_of_ne_zero hb) (h b)) (b.lt_pow_self ha_gt_one), fun h => by cases h <;> simp [*]⟩ alias ⟨_, Dvd.multiplicity_pos⟩ := dvd_iff_multiplicity_pos diff --git a/Mathlib/RingTheory/WittVector/Frobenius.lean b/Mathlib/RingTheory/WittVector/Frobenius.lean index 79cc7194f5851..8ae9671d90379 100644 --- a/Mathlib/RingTheory/WittVector/Frobenius.lean +++ b/Mathlib/RingTheory/WittVector/Frobenius.lean @@ -124,7 +124,7 @@ theorem map_frobeniusPoly.key₂ {n i j : ℕ} (hi : i ≤ n) (hj : j < p ^ (n - tsub_add_cancel_of_le (le_tsub_of_add_le_right ((le_tsub_iff_left hi).mp h₁))] have hle : p ^ m ≤ j + 1 := h ▸ Nat.le_of_dvd j.succ_pos (pow_multiplicity_dvd _ _) exact ⟨(Nat.pow_le_pow_iff_right hp.1.one_lt).1 (hle.trans hj), - Nat.le_of_lt_succ ((Nat.lt_pow_self hp.1.one_lt m).trans_le hle)⟩ + Nat.le_of_lt_succ ((m.lt_pow_self hp.1.one_lt).trans_le hle)⟩ theorem map_frobeniusPoly (n : ℕ) : MvPolynomial.map (Int.castRingHom ℚ) (frobeniusPoly p n) = frobeniusPolyRat p n := by diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index 4229049cf7321..d2fe0b2b85cf9 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -73,7 +73,7 @@ instance [LocallyCompactSpace H] : LocallyCompactSpace (PontryaginDual H) := by refine lt_of_lt_of_le ht ?_ rw [div_le_iff₀' (pow_pos two_pos _), ← div_le_iff₀ hx] refine (Nat.le_ceil (Real.pi / x)).trans ?_ - exact_mod_cast (Nat.le_succ _).trans (Nat.lt_two_pow _).le + exact_mod_cast (Nat.le_succ _).trans Nat.lt_two_pow.le variable {A B C G} From 386ab515e2f0ed7be0211cbf37b1cd40b8cb4321 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 03:22:16 +0000 Subject: [PATCH 112/250] chore(Multilinear/Basic): address a porting note (#19443) --- Mathlib/LinearAlgebra/Multilinear/Basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/LinearAlgebra/Multilinear/Basic.lean b/Mathlib/LinearAlgebra/Multilinear/Basic.lean index 96de9639e3cfd..d12337d2b24a3 100644 --- a/Mathlib/LinearAlgebra/Multilinear/Basic.lean +++ b/Mathlib/LinearAlgebra/Multilinear/Basic.lean @@ -285,7 +285,6 @@ def constOfIsEmpty [IsEmpty ι] (m : M₂) : MultilinearMap R M₁ M₂ where end --- Porting note: Included `DFunLike.coe` to avoid strange CoeFun instance for Equiv /-- Given a multilinear map `f` on `n` variables (parameterized by `Fin n`) and a subset `s` of `k` of these variables, one gets a new multilinear map on `Fin k` by varying these variables, and fixing the other ones equal to a given value `z`. It is denoted by `f.restr s hk z`, where `hk` is a @@ -293,7 +292,7 @@ proof that the cardinality of `s` is `k`. The implicit identification between `F we use is the canonical (increasing) bijection. -/ def restr {k n : ℕ} (f : MultilinearMap R (fun _ : Fin n => M') M₂) (s : Finset (Fin n)) (hk : #s = k) (z : M') : MultilinearMap R (fun _ : Fin k => M') M₂ where - toFun v := f fun j => if h : j ∈ s then v ((DFunLike.coe (s.orderIsoOfFin hk).symm) ⟨j, h⟩) else z + toFun v := f fun j => if h : j ∈ s then v ((s.orderIsoOfFin hk).symm ⟨j, h⟩) else z /- Porting note: The proofs of the following two lemmas used to only use `erw` followed by `simp`, but it seems `erw` no longer unfolds or unifies well enough to work without more help. -/ map_update_add' v i x y := by From 04044ff3f96e5c4a520922b00450ef0dbef6e4e8 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 03:43:22 +0000 Subject: [PATCH 113/250] chore: move pointwise set files to `Algebra._.Pointwise` (#19002) --- Mathlib.lean | 8 ++++---- Mathlib/Algebra/Algebra/Operations.lean | 2 +- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 4 ++-- Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean | 2 +- .../Group/Pointwise/Set}/BigOperators.lean | 0 Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 2 +- .../Pointwise => Algebra/Group/Pointwise/Set}/Finite.lean | 0 .../Group/Pointwise/Set}/ListOfFn.lean | 0 .../Set => Algebra/Order/Group}/Pointwise/Interval.lean | 0 Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++-- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/GroupTheory/Finiteness.lean | 2 +- Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean | 2 +- Mathlib/Topology/Algebra/Field.lean | 2 +- Mathlib/Topology/Algebra/Order/Field.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Real.lean | 2 +- 16 files changed, 17 insertions(+), 17 deletions(-) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/BigOperators.lean (100%) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/Finite.lean (100%) rename Mathlib/{Data/Set/Pointwise => Algebra/Group/Pointwise/Set}/ListOfFn.lean (100%) rename Mathlib/{Data/Set => Algebra/Order/Group}/Pointwise/Interval.lean (100%) diff --git a/Mathlib.lean b/Mathlib.lean index 58d4cfdeb00c9..f3b4520b86d56 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -294,7 +294,10 @@ import Mathlib.Algebra.Group.Pi.Lemmas import Mathlib.Algebra.Group.Pointwise.Finset.Basic import Mathlib.Algebra.Group.Pointwise.Finset.Interval import Mathlib.Algebra.Group.Pointwise.Set.Basic +import Mathlib.Algebra.Group.Pointwise.Set.BigOperators import Mathlib.Algebra.Group.Pointwise.Set.Card +import Mathlib.Algebra.Group.Pointwise.Set.Finite +import Mathlib.Algebra.Group.Pointwise.Set.ListOfFn import Mathlib.Algebra.Group.Prod import Mathlib.Algebra.Group.Semiconj.Basic import Mathlib.Algebra.Group.Semiconj.Defs @@ -661,6 +664,7 @@ import Mathlib.Algebra.Order.Group.OrderIso import Mathlib.Algebra.Order.Group.PiLex import Mathlib.Algebra.Order.Group.Pointwise.Bounds import Mathlib.Algebra.Order.Group.Pointwise.CompleteLattice +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Algebra.Order.Group.PosPart import Mathlib.Algebra.Order.Group.Prod import Mathlib.Algebra.Order.Group.Synonym @@ -2820,11 +2824,7 @@ import Mathlib.Data.Set.Operations import Mathlib.Data.Set.Opposite import Mathlib.Data.Set.Pairwise.Basic import Mathlib.Data.Set.Pairwise.Lattice -import Mathlib.Data.Set.Pointwise.BigOperators -import Mathlib.Data.Set.Pointwise.Finite -import Mathlib.Data.Set.Pointwise.Interval import Mathlib.Data.Set.Pointwise.Iterate -import Mathlib.Data.Set.Pointwise.ListOfFn import Mathlib.Data.Set.Pointwise.SMul import Mathlib.Data.Set.Pointwise.Support import Mathlib.Data.Set.Prod diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index de1a4621d5777..960f805fd05b7 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -6,9 +6,9 @@ Authors: Kenny Lau import Mathlib.Algebra.Algebra.Bilinear import Mathlib.Algebra.Algebra.Opposite import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Algebra.Group.Pointwise.Set.BigOperators import Mathlib.Algebra.GroupWithZero.NonZeroDivisors import Mathlib.Algebra.Module.Submodule.Pointwise -import Mathlib.Data.Set.Pointwise.BigOperators import Mathlib.Data.Set.Semiring import Mathlib.GroupTheory.GroupAction.SubMulAction.Pointwise diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index a7cd803e05c07..4330e3cf9abae 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -5,11 +5,11 @@ Authors: Floris van Doorn, Yaël Dillies -/ import Mathlib.Algebra.BigOperators.Group.Finset import Mathlib.Algebra.Group.Action.Pi +import Mathlib.Algebra.Group.Pointwise.Set.Finite +import Mathlib.Algebra.Group.Pointwise.Set.ListOfFn import Mathlib.Data.Finset.Density import Mathlib.Data.Finset.Max import Mathlib.Data.Finset.NAry -import Mathlib.Data.Set.Pointwise.Finite -import Mathlib.Data.Set.Pointwise.ListOfFn import Mathlib.Data.Set.Pointwise.SMul /-! diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean index 479121a3d4a1b..d06c3e7259b13 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Interval.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Order.Interval.Finset.Defs /-! # Pointwise operations on intervals diff --git a/Mathlib/Data/Set/Pointwise/BigOperators.lean b/Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/BigOperators.lean rename to Mathlib/Algebra/Group/Pointwise/Set/BigOperators.lean diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 06a3718af24d0..9472500e7fcc9 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -3,8 +3,8 @@ 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.Group.Pointwise.Set.Finite import Mathlib.Data.Set.Card -import Mathlib.Data.Set.Pointwise.Finite import Mathlib.SetTheory.Cardinal.Finite /-! diff --git a/Mathlib/Data/Set/Pointwise/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/Finite.lean rename to Mathlib/Algebra/Group/Pointwise/Set/Finite.lean diff --git a/Mathlib/Data/Set/Pointwise/ListOfFn.lean b/Mathlib/Algebra/Group/Pointwise/Set/ListOfFn.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/ListOfFn.lean rename to Mathlib/Algebra/Group/Pointwise/Set/ListOfFn.lean diff --git a/Mathlib/Data/Set/Pointwise/Interval.lean b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean similarity index 100% rename from Mathlib/Data/Set/Pointwise/Interval.lean rename to Mathlib/Algebra/Order/Group/Pointwise/Interval.lean diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index a81ab41393a9d..eb3c22b61051f 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -5,10 +5,10 @@ Authors: Patrick Massot, Johannes Hölzl -/ import Mathlib.Algebra.Algebra.NonUnitalSubalgebra import Mathlib.Algebra.Algebra.Subalgebra.Basic +import Mathlib.Algebra.Field.Subfield.Defs +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Analysis.Normed.Group.Constructions import Mathlib.Analysis.Normed.Group.Submodule -import Mathlib.Data.Set.Pointwise.Interval -import Mathlib.Algebra.Field.Subfield.Defs /-! # Normed fields diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index b0a7e99a26fb2..b41af8e41d56f 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -3,9 +3,9 @@ Copyright (c) 2019 Floris van Doorn. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn -/ +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Analysis.SpecificLimits.Basic import Mathlib.Data.Rat.Cardinal -import Mathlib.Data.Set.Pointwise.Interval import Mathlib.SetTheory.Cardinal.Continuum /-! diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index 80895c9fc92ee..bd31d6b8743b4 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -3,8 +3,8 @@ Copyright (c) 2021 Riccardo Brasca. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Riccardo Brasca -/ +import Mathlib.Algebra.Group.Pointwise.Set.Finite import Mathlib.Algebra.Group.Subgroup.Pointwise -import Mathlib.Data.Set.Pointwise.Finite import Mathlib.GroupTheory.QuotientGroup.Defs import Mathlib.SetTheory.Cardinal.Finite diff --git a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean index 71f99e5c85dde..259da107ba011 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/AffineMap.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.LinearAlgebra.AffineSpace.Basic import Mathlib.LinearAlgebra.BilinearMap import Mathlib.LinearAlgebra.Pi diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index 7e78ef983d7f6..cc34c6223561e 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -5,10 +5,10 @@ Authors: Patrick Massot, Kim Morrison -/ import Mathlib.Algebra.Field.Subfield.Basic import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Topology.Algebra.GroupWithZero import Mathlib.Topology.Algebra.Ring.Basic import Mathlib.Topology.Order.LocalExtr -import Mathlib.Data.Set.Pointwise.Interval /-! # Topological fields diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index 51f7585e7e00c..ee7ea3caa6442 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Benjamin Davidson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Benjamin Davidson, Devon Tuma, Eric Rodriguez, Oliver Nash -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Order.Filter.AtTopBot.Field import Mathlib.Topology.Algebra.Field import Mathlib.Topology.Algebra.Order.Group diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean index ae477178577c5..baa4f534f75cc 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean @@ -3,7 +3,7 @@ Copyright (c) 2015, 2017 Jeremy Avigad. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Robert Y. Lewis, Johannes Hölzl, Mario Carneiro, Sébastien Gouëzel -/ -import Mathlib.Data.Set.Pointwise.Interval +import Mathlib.Algebra.Order.Group.Pointwise.Interval import Mathlib.Topology.MetricSpace.Pseudo.Pi /-! From aa050b47ffec260eb307d5ad30c2cc02a0a8c5d4 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 06:35:02 +0000 Subject: [PATCH 114/250] feat(bench_summary): report no significant changes and job_id (#18812) If the `!bench` summary contains no entry with a difference in instructions of at least 10^9, then the bot posts a message confirming this information, rather than silently failing. The bot also posts a link to the workflow CI run that generated the message. --- .github/workflows/bench_summary_comment.yml | 2 +- scripts/bench_summary.lean | 16 ++++++++++++---- 2 files changed, 13 insertions(+), 5 deletions(-) diff --git a/.github/workflows/bench_summary_comment.yml b/.github/workflows/bench_summary_comment.yml index 6672fb8443e98..fe14ecdc32a79 100644 --- a/.github/workflows/bench_summary_comment.yml +++ b/.github/workflows/bench_summary_comment.yml @@ -27,7 +27,7 @@ jobs: run: | { cat scripts/bench_summary.lean - printf $'run_cmd BenchAction.addBenchSummaryComment %s "leanprover-community/mathlib4"' "${PR}" + printf $'run_cmd BenchAction.addBenchSummaryComment %s "leanprover-community/mathlib4" %s' "${PR}" "${{ github.run_id }}" } | lake env lean --stdin env: diff --git a/scripts/bench_summary.lean b/scripts/bench_summary.lean index 9c9407ea066f4..17f0467b2ef97 100644 --- a/scripts/bench_summary.lean +++ b/scripts/bench_summary.lean @@ -147,6 +147,7 @@ open Lean Elab Command in as a comment to a pull request. It takes as input * the number `PR` and the name `repo` as a `String` containing the relevant pull-request (it reads and posts comments there) +* the optional `jobID` string for reporting the action that produced the output * the `String` `tempFile` of a temporary file where the command stores transient information. The code itself interfaces with the shell to retrieve and process json data and eventually @@ -162,9 +163,10 @@ Here is a summary of the steps: * process the final string to produce a summary (using `benchOutput`), * finally post the resulting output to the PR (using `gh pr comment ...`). -/ -def addBenchSummaryComment (PR : Nat) (repo : String) +def addBenchSummaryComment (PR : Nat) (repo : String) (jobID : String := "") (author : String := "leanprover-bot") (tempFile : String := "benchOutput.json") : CommandElabM Unit := do + let job_msg := s!"\n[CI run](https://github.com/{repo}/actions/runs/{jobID})" let PR := s!"{PR}" let jq := s!".comments | last | select(.author.login==\"{author}\") | .body" @@ -197,10 +199,10 @@ def addBenchSummaryComment (PR : Nat) (repo : String) IO.FS.writeFile (tempFile ++ ".src") bench -- Extract all instruction changes whose magnitude is larger than `threshold`. - let threshold := s!"{10 ^ 9}" + let threshold := 10 ^ 9 let jq1 : IO.Process.SpawnArgs := { cmd := "jq" - args := #["-r", "--arg", "thr", threshold, + args := #["-r", "--arg", "thr", s!"{threshold}", ".differences | .[] | ($thr|tonumber) as $th | select(.dimension.metric == \"instructions\" and ((.diff >= $th) or (.diff <= -$th)))", (tempFile ++ ".src")] } @@ -220,6 +222,12 @@ def addBenchSummaryComment (PR : Nat) (repo : String) jq -c '[\{file: .dimension.benchmark, diff: .diff, reldiff: .reldiff}]' {tempFile} > \ {tempFile}.2" let secondFilter ← IO.Process.run jq2 + if secondFilter == "" then + let _ ← IO.Process.run + { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", + s!"No benchmark entry differed by at least {formatDiff threshold} instructions." ++ + job_msg] } + else IO.FS.writeFile tempFile secondFilter let jq3 : IO.Process.SpawnArgs := { cmd := "jq", args := #["-n", "reduce inputs as $in (null; . + $in)", tempFile] } @@ -230,7 +238,7 @@ def addBenchSummaryComment (PR : Nat) (repo : String) IO.println report -- Post the computed summary as a github comment. let add_comment : IO.Process.SpawnArgs := - { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", report] } + { cmd := "gh", args := #["pr", "comment", PR, "--repo", repo, "--body", report ++ job_msg] } let _ ← IO.Process.run add_comment end BenchAction From aad267d3a84144ad19fe4ea77f9ad0fd80679f95 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 07:23:24 +0000 Subject: [PATCH 115/250] chore(discover-lean-pr-testing): ignore stage0 PRs, more tracing (#19405) --- .github/workflows/discover-lean-pr-testing.yml | 18 +++++++++++++----- 1 file changed, 13 insertions(+), 5 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 8210a3e66daab..f9df1791979d7 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -36,8 +36,10 @@ jobs: # The "./" in front of the path is important for `git show` OLD=$(git show "$PREV_COMMIT":./lean-toolchain | cut -f2 -d:) - echo "new=$NEW" >> "$GITHUB_OUTPUT" + echo "old=$OLD" + echo "new=$NEW" echo "old=$OLD" >> "$GITHUB_OUTPUT" + echo "new=$NEW" >> "$GITHUB_OUTPUT" - name: Clone lean4 repository and get PRs id: get-prs @@ -63,7 +65,10 @@ jobs: # This will only fetch commits newer than previously fetched commits (ie $OLD) git fetch origin tag "$NEW" --no-tags - PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/') + # Find all commits to lean4 between the $OLD and $NEW toolchains + # and extract the github PR numbers + # (drop anything that doesn't look like a PR number from the final result) + PRS=$(git log --oneline "$OLD..$NEW" | sed 's/.*(#\([0-9]\+\))$/\1/' | grep -E '^[0-9]+$') # Output the PRs echo "$PRS" @@ -74,12 +79,14 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" + echo "=== $PRS ========================" echo "$PRS" - echo "====================" echo "$PRS" | tr ' ' '\n' > prs.txt + echo "=== prs.txt =====================" + cat prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) + echo "=== $MATCHING_BRANCHES ==========" echo "$MATCHING_BRANCHES" - echo "====================" # Initialize an empty variable to store branches with relevant diffs RELEVANT_BRANCHES="" @@ -92,11 +99,12 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES\n- $BRANCH" + RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$BRANCH" fi done # Output the relevant branches + echo "=== $RELEVANT_BRANCHES ==========" echo "'$RELEVANT_BRANCHES'" printf "branches<> "$GITHUB_OUTPUT" From 55a97c9bd9fc11ce7a0cec05b9b355ffcbaa63ba Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 07:56:44 +0000 Subject: [PATCH 116/250] chore(discover-lean-pr-testing): fix escaping hazard (#19448) --- .github/workflows/discover-lean-pr-testing.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index f9df1791979d7..02e087dfb6085 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -79,13 +79,13 @@ jobs: run: | # Use the PRs from the previous step PRS="${{ steps.get-prs.outputs.prs }}" - echo "=== $PRS ========================" + echo "=== PRS =========================" echo "$PRS" echo "$PRS" | tr ' ' '\n' > prs.txt echo "=== prs.txt =====================" cat prs.txt MATCHING_BRANCHES=$(git branch -r | grep -f prs.txt) - echo "=== $MATCHING_BRANCHES ==========" + echo "=== MATCHING_BRANCHES ===========" echo "$MATCHING_BRANCHES" # Initialize an empty variable to store branches with relevant diffs @@ -104,7 +104,7 @@ jobs: done # Output the relevant branches - echo "=== $RELEVANT_BRANCHES ==========" + echo "=== RELEVANT_BRANCHES ===========" echo "'$RELEVANT_BRANCHES'" printf "branches<> "$GITHUB_OUTPUT" From a41d8f43ee87309b9bb3c4f23c473a5fcbc2e922 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 08:25:35 +0000 Subject: [PATCH 117/250] feat: trigger automated Zulip emojis via `bors x` command (#19371) Changes the way in which emoji reactions to `delegated`, `ready-to-merge` and `merged` are applied in Zulip: instead of running periodically a script to add the reactions, the reactions are now applied at the same time that the label is applied, or when PRs are merged into master. --- .github/workflows/maintainer_bors.yml | 28 ++++- .../workflows/zulip_emoji_merge_delegate.yaml | 21 ++-- scripts/README.md | 9 +- scripts/zulip_emoji_merge_delegate.py | 105 +++++++++++------- 4 files changed, 109 insertions(+), 54 deletions(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 8b06a44a1b926..290ea7c66fa1b 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -84,4 +84,30 @@ jobs: curl --request DELETE \ --url "https://api.github.com/repos/${{ github.repository }}/issues/${{ github.event.issue.number }}${{ github.event.pull_request.number }}/labels/awaiting-author" \ --header 'authorization: Bearer ${{ secrets.TRIAGE_TOKEN }}' - # comment uses ${{ secrets.GITHUB_TOKEN }} + + - name: Set up Python + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + uses: actions/setup-python@v5 + with: + python-version: '3.x' + + - name: Install dependencies + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + run: | + python -m pip install --upgrade pip + pip install zulip + + - name: Run Zulip Emoji Merge Delegate Script + if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && + ( steps.user_permission.outputs.require-result == 'true' || + steps.merge_or_delegate.outputs.bot == 'true' ) }} + env: + ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} + ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com + ZULIP_SITE: https://leanprover.zulipchat.com + run: | + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "${{ steps.merge_or_delegate.outputs.mOrD }}" "${{ github.event.issue.number }}${{ github.event.pull_request.number }}" diff --git a/.github/workflows/zulip_emoji_merge_delegate.yaml b/.github/workflows/zulip_emoji_merge_delegate.yaml index e44b4d07004d9..3de0bc2fd1aa4 100644 --- a/.github/workflows/zulip_emoji_merge_delegate.yaml +++ b/.github/workflows/zulip_emoji_merge_delegate.yaml @@ -1,19 +1,19 @@ name: Zulip Emoji Merge Delegate on: - schedule: - - cron: '0 * * * *' # Runs every hour + push: + branches: + - master jobs: - zulip-emoji-merge-delegate: + zulip-emoji-merged: runs-on: ubuntu-latest steps: - name: Checkout mathlib repository uses: actions/checkout@v4 with: - sparse-checkout: | - scripts + fetch-depth: 0 # donwload the full repository - name: Set up Python uses: actions/setup-python@v5 @@ -30,6 +30,13 @@ jobs: ZULIP_API_KEY: ${{ secrets.ZULIP_API_KEY }} ZULIP_EMAIL: github-mathlib4-bot@leanprover.zulipchat.com ZULIP_SITE: https://leanprover.zulipchat.com - GITHUB_TOKEN: ${{ secrets.GITHUB_TOKEN }} run: | - python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "$GITHUB_TOKEN" + # scan the commits of the past 10 minutes, assuming that the timezone of the current machine + # is the same that performed the commit + git log --since="10 minutes ago" --pretty=oneline + + printf $'Scanning commits:\n%s\n\nContinuing\n' "$(git log --since="10 minutes ago" --pretty=oneline)" + while read -r pr_number; do + printf $'Running the python script with pr "%s"\n' "${pr_number}" + python scripts/zulip_emoji_merge_delegate.py "$ZULIP_API_KEY" "$ZULIP_EMAIL" "$ZULIP_SITE" "[Merged by Bors]" "${pr_number}" + done < <(git log --oneline -n 10 | awk -F# '{ gsub(/)$/, ""); print $NF}') diff --git a/scripts/README.md b/scripts/README.md index 44156380d5e62..3962da0193667 100644 --- a/scripts/README.md +++ b/scripts/README.md @@ -93,9 +93,12 @@ please do not add new entries to these files. PRs removing (the need for) entrie to the appropriate topic on zulip. - `count-trans-deps.py`, `import-graph-report.py` and `import_trans_difference.sh` produce various summaries of changes in transitive imports that the `PR_summary` message incorporates. -- `zulip_emoji_merge_delegate.py` is called every hour by a Github action cronjob. - It looks through the latest 200 zulip posts: if a message mentions a PR that is delegated, or sent to bors, or merged, - then this script will post an emoji reaction `:peace_sign:`, or `:bors:`, or `:merge:` respectively to the message. +- `zulip_emoji_merge_delegate.py` is called + * every time a `bors d`, `bors merge` or `bors r+` comment is added to a PR and + * on every push to `master`. + It looks through all zulip posts containing a reference to the relevant PR + (delegated, or sent to bors, or merged) and it will post an emoji reaction + `:peace_sign:`, or `:bors:`, or `:merge:` respectively to the message. - `late_importers.sh` is the main script used by the `latest_import.yml` action: it formats the `linter.minImports` output, summarizing the data in a table. See the module docs of `late_importers.sh` for further details. diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index d715df3bfbf9a..b279fbc6e1e11 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -1,17 +1,20 @@ #!/usr/bin/env python3 import sys import zulip -import requests import re # Usage: -# python scripts/zulip_emoji_merge_delegate.py $ZULIP_API_KEY $ZULIP_EMAIL $ZULIP_SITE $GITHUB_TOKEN +# python scripts/zulip_emoji_merge_delegate.py $ZULIP_API_KEY $ZULIP_EMAIL $ZULIP_SITE $LABEL $PR_NUMBER # See .github/workflows/zulip_emoji_merge_delegate.yaml for the meaning of these variables ZULIP_API_KEY = sys.argv[1] ZULIP_EMAIL = sys.argv[2] ZULIP_SITE = sys.argv[3] -GITHUB_TOKEN = sys.argv[4] +LABEL = sys.argv[4] +PR_NUMBER = sys.argv[5] + +print(f"LABEL: '{LABEL}'") +print(f"PR_NUMBER: '{PR_NUMBER}'") # Initialize Zulip client client = zulip.Client( @@ -22,58 +25,74 @@ # Fetch the last 200 messages response = client.get_messages({ - "anchor": "newest", - "num_before": 200, - "num_after": 0, - "narrow": [{"operator": "channel", "operand": "PR reviews"}], + "operator": "search", + "operand": f"https://github\.com/leanprover-community/mathlib4/pull/{PR_NUMBER}", }) messages = response['messages'] -pr_pattern = re.compile(r'https://github\.com/leanprover-community/mathlib4/pull/(\d+)') for message in messages: content = message['content'] - match = pr_pattern.search(content) - if match: - pr_number = match.group(1) - # Check for emoji reactions - reactions = message['reactions'] - has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) - has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) - has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + print(f"matched: '{message}'") + # Check for emoji reactions + reactions = message['reactions'] + has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) + has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) + has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + + pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{PR_NUMBER}" + + print('Removing peace_sign') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + print(f"result: '{result}'") + print('Removing bors') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "bors" + }) + print(f"result: '{result}'") + + print('Removing merge') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) + print(f"result: '{result}'") - pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{pr_number}" - pr_response = requests.get(pr_url, headers={"Authorization": GITHUB_TOKEN}) - pr_data = pr_response.json() - labels = [label['name'] for label in pr_data['labels']] + if 'delegated' == LABEL: + print('adding delegated') - if 'delegated' in labels: - client.add_reaction({ + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + elif 'ready-to-merge' == LABEL: + print('adding ready-to-merge') + if has_peace_sign: + client.remove_reaction({ "message_id": message['id'], "emoji_name": "peace_sign" }) - elif 'ready-to-merge' in labels: - if has_peace_sign: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - client.add_reaction({ + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "bors" + }) + elif LABEL.startswith("[Merged by Bors]"): + print('adding [Merged by Bors]') + if has_peace_sign: + client.remove_reaction({ "message_id": message['id'], - "emoji_name": "bors" + "emoji_name": "peace_sign" }) - elif pr_data['title'].startswith("[Merged by Bors]"): - if has_peace_sign: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - if has_bors: - client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - client.add_reaction({ + if has_bors: + client.remove_reaction({ "message_id": message['id'], - "emoji_name": "merge" + "emoji_name": "bors" }) + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) From 27f81bc0f58a91ff043b663f3277e9bb11d89ea3 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Mon, 25 Nov 2024 08:37:46 +0000 Subject: [PATCH 118/250] chore: bump to nightly-2024-11-25 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 22be24ab2e033..1e6741ae1e185 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-24 +leanprover/lean4:nightly-2024-11-25 From 4eab2d4746acfc4e831ac0bac6d3440679c55f5c Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Mon, 25 Nov 2024 08:40:48 +0000 Subject: [PATCH 119/250] feat: the free locus of a (finitely presented) module (#18691) Co-authored-by: Andrew Yang <36414270+erdOne@users.noreply.github.com> --- Mathlib.lean | 1 + .../Algebra/Module/FinitePresentation.lean | 35 ++- Mathlib/Algebra/Module/FreeLocus.lean | 227 ++++++++++++++++++ .../Algebra/Module/LocalizedModule/Basic.lean | 47 +++- Mathlib/LinearAlgebra/FreeModule/Basic.lean | 21 ++ Mathlib/RingTheory/Finiteness/Basic.lean | 11 + .../LocalProperties/Projective.lean | 43 +++- 7 files changed, 379 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/Module/FreeLocus.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3b4520b86d56..7e32e8e68409c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -518,6 +518,7 @@ import Mathlib.Algebra.Module.Equiv.Basic import Mathlib.Algebra.Module.Equiv.Defs import Mathlib.Algebra.Module.Equiv.Opposite import Mathlib.Algebra.Module.FinitePresentation +import Mathlib.Algebra.Module.FreeLocus import Mathlib.Algebra.Module.GradedModule import Mathlib.Algebra.Module.Hom import Mathlib.Algebra.Module.Injective diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index a4d6e21f87576..553e1aee1ced2 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -4,10 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ import Mathlib.LinearAlgebra.FreeModule.Finite.Basic +import Mathlib.LinearAlgebra.TensorProduct.RightExactness import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Finiteness.Projective -import Mathlib.RingTheory.Localization.Module +import Mathlib.RingTheory.Finiteness.TensorProduct +import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Noetherian.Defs + /-! # Finitely Presented Modules @@ -232,6 +235,36 @@ section CommRing variable {R M N N'} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) (f : N →ₗ[R] N') [IsLocalizedModule S f] +open TensorProduct in +instance {A} [CommRing A] [Algebra R A] [Module.FinitePresentation R M] : + Module.FinitePresentation A (A ⊗[R] M) := by + classical + obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R M + have inst := Module.finitePresentation_of_projective A (A ⊗[R] (Fin n → R)) + apply Module.finitePresentation_of_surjective (f.baseChange A) + (LinearMap.lTensor_surjective A hf) + have : Function.Exact ((LinearMap.ker f).subtype.baseChange A) (f.baseChange A) := + lTensor_exact A f.exact_subtype_ker_map hf + rw [LinearMap.exact_iff] at this + rw [this, ← Submodule.map_top] + apply Submodule.FG.map + have : Module.Finite R (LinearMap.ker f) := + ⟨(Submodule.fg_top _).mpr (Module.FinitePresentation.fg_ker f hf)⟩ + exact Module.Finite.out (R := A) (M := A ⊗[R] LinearMap.ker f) + +open TensorProduct in +lemma FinitePresentation.of_isBaseChange + {A} [CommRing A] [Algebra R A] [Module A N] [IsScalarTower R A N] + (f : M →ₗ[R] N) (h : IsBaseChange A f) [Module.FinitePresentation R M] : + Module.FinitePresentation A N := + Module.finitePresentation_of_surjective + h.equiv.toLinearMap h.equiv.surjective (by simpa using Submodule.fg_bot) + +open TensorProduct in +instance (S : Submonoid R) [Module.FinitePresentation R M] : + Module.FinitePresentation (Localization S) (LocalizedModule S M) := + FinitePresentation.of_isBaseChange (LocalizedModule.mkLinearMap S M) + ((isLocalizedModule_iff_isBaseChange S _ _).mp inferInstance) lemma Module.FinitePresentation.exists_lift_of_isLocalizedModule [h : Module.FinitePresentation R M] (g : M →ₗ[R] N') : diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean new file mode 100644 index 0000000000000..9f42786132f02 --- /dev/null +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -0,0 +1,227 @@ +/- +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.PrimeSpectrum.Basic +import Mathlib.RingTheory.Flat.Stability +import Mathlib.RingTheory.LocalProperties.Projective +import Mathlib.RingTheory.LocalRing.Module +import Mathlib.RingTheory.Localization.Free +import Mathlib.RingTheory.Localization.LocalizationLocalization +import Mathlib.Topology.LocallyConstant.Basic + +/-! + +# The free locus of a module + +## Main definitions and results + +Let `M` be a finitely presented `R`-module. +- `Module.freeLocus`: The set of points `x` in `Spec R` such that `Mₓ` is free over `Rₓ`. +- `Module.freeLocus_eq_univ_iff`: + The free locus is the whole `Spec R` if and only if `M` is projective. +- `Module.basicOpen_subset_freeLocus_iff`: `D(f)` is contained in the free locus if and only if + `M_f` is projective over `R_f`. +- `Module.rankAtStalk`: The function `Spec R → ℕ` sending `x` to `rank_{Rₓ} Mₓ`. +- `Module.isLocallyConstant_rankAtStalk`: + If `M` is flat over `R`, then `rankAtStalk` is locally constant. + +-/ + +universe uR uM + +variable (R : Type uR) (M : Type uM) [CommRing R] [AddCommGroup M] [Module R M] + +namespace Module + +open PrimeSpectrum TensorProduct + +/-- The free locus of a module, i.e. the set of primes `p` such that `Mₚ` is free over `Rₚ`. -/ +def freeLocus : Set (PrimeSpectrum R) := + { p | Module.Free (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) } + +variable {R M} + +lemma mem_freeLocus {p} : p ∈ freeLocus R M ↔ + Module.Free (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) := + Iff.rfl + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma mem_freeLocus_of_isLocalization (p : PrimeSpectrum R) + (Rₚ Mₚ) [CommRing Rₚ] [Algebra R Rₚ] [IsLocalization.AtPrime Rₚ p.asIdeal] + [AddCommGroup Mₚ] [Module R Mₚ] (f : M →ₗ[R] Mₚ) [IsLocalizedModule p.asIdeal.primeCompl f] + [Module Rₚ Mₚ] [IsScalarTower R Rₚ Mₚ] : + p ∈ freeLocus R M ↔ Module.Free Rₚ Mₚ := by + apply Module.Free.iff_of_ringEquiv (IsLocalization.algEquiv p.asIdeal.primeCompl + (Localization.AtPrime p.asIdeal) Rₚ).toRingEquiv + refine { __ := IsLocalizedModule.iso p.asIdeal.primeCompl f, map_smul' := ?_ } + intro r x + obtain ⟨r, s, rfl⟩ := IsLocalization.mk'_surjective p.asIdeal.primeCompl r + apply ((Module.End_isUnit_iff _).mp (IsLocalizedModule.map_units f s)).1 + simp only [AddHom.toFun_eq_coe, LinearMap.coe_toAddHom, LinearEquiv.coe_coe, + algebraMap_end_apply, AlgEquiv.toRingEquiv_eq_coe, + AlgEquiv.toRingEquiv_toRingHom, RingHom.coe_coe, IsLocalization.algEquiv_apply, + IsLocalization.map_id_mk'] + simp only [← map_smul, ← smul_assoc, IsLocalization.smul_mk'_self, algebraMap_smul, + IsLocalization.map_id_mk'] + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma mem_freeLocus_iff_tensor (p : PrimeSpectrum R) + (Rₚ) [CommRing Rₚ] [Algebra R Rₚ] [IsLocalization.AtPrime Rₚ p.asIdeal] : + p ∈ freeLocus R M ↔ Module.Free Rₚ (Rₚ ⊗[R] M) := by + have := (isLocalizedModule_iff_isBaseChange p.asIdeal.primeCompl _ _).mpr + (TensorProduct.isBaseChange R M Rₚ) + exact mem_freeLocus_of_isLocalization p Rₚ (f := TensorProduct.mk R Rₚ M 1) + +lemma freeLocus_congr {M'} [AddCommGroup M'] [Module R M'] (e : M ≃ₗ[R] M') : + freeLocus R M = freeLocus R M' := by + ext p + exact mem_freeLocus_of_isLocalization _ _ _ + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M' ∘ₗ e.toLinearMap) + +open TensorProduct in +lemma comap_freeLocus_le {A} [CommRing A] [Algebra R A] : + comap (algebraMap R A) ⁻¹' freeLocus R M ≤ freeLocus A (A ⊗[R] M) := by + intro p hp + let Rₚ := Localization.AtPrime (comap (algebraMap R A) p).asIdeal + let Aₚ := Localization.AtPrime p.asIdeal + rw [Set.mem_preimage, mem_freeLocus_iff_tensor _ Rₚ] at hp + rw [mem_freeLocus_iff_tensor _ Aₚ] + letI : Algebra Rₚ Aₚ := (Localization.localRingHom + (comap (algebraMap R A) p).asIdeal p.asIdeal (algebraMap R A) rfl).toAlgebra + have : IsScalarTower R Rₚ Aₚ := IsScalarTower.of_algebraMap_eq' + (by simp [RingHom.algebraMap_toAlgebra, Localization.localRingHom, + ← IsScalarTower.algebraMap_eq]) + let e := AlgebraTensorModule.cancelBaseChange R Rₚ Aₚ Aₚ M ≪≫ₗ + (AlgebraTensorModule.cancelBaseChange R A Aₚ Aₚ M).symm + exact .of_equiv e + +lemma freeLocus_localization (S : Submonoid R) : + freeLocus (Localization S) (LocalizedModule S M) = + comap (algebraMap R _) ⁻¹' freeLocus R M := by + ext p + simp only [Set.mem_preimage] + let p' := p.asIdeal.comap (algebraMap R _) + have hp' : S ≤ p'.primeCompl := fun x hx H ↦ + p.isPrime.ne_top (Ideal.eq_top_of_isUnit_mem _ H (IsLocalization.map_units _ ⟨x, hx⟩)) + let Rₚ := Localization.AtPrime p' + let Mₚ := LocalizedModule p'.primeCompl M + letI : Algebra (Localization S) Rₚ := + IsLocalization.localizationAlgebraOfSubmonoidLe _ _ S p'.primeCompl hp' + have : IsScalarTower R (Localization S) Rₚ := + IsLocalization.localization_isScalarTower_of_submonoid_le .. + have : IsLocalization.AtPrime Rₚ p.asIdeal := by + have := IsLocalization.isLocalization_of_submonoid_le (Localization S) Rₚ _ _ hp' + apply IsLocalization.isLocalization_of_is_exists_mul_mem _ + (Submonoid.map (algebraMap R (Localization S)) p'.primeCompl) + · rintro _ ⟨x, hx, rfl⟩; exact hx + · rintro ⟨x, hx⟩ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective S x + refine ⟨algebraMap _ _ s.1, x, fun H ↦ hx ?_, by simp⟩ + rw [IsLocalization.mk'_eq_mul_mk'_one] + exact Ideal.mul_mem_right _ _ H + letI : Module (Localization S) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) + have : IsScalarTower R (Localization S) Mₚ := + ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by + simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + have : IsScalarTower (Localization S) Rₚ Mₚ := + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) + (LocalizedModule.mkLinearMap p'.primeCompl M)).extendScalarsOfIsLocalization S + (Localization S) + have : IsLocalizedModule p.asIdeal.primeCompl l := by + have : IsLocalizedModule p'.primeCompl (l.restrictScalars R) := + inferInstanceAs (IsLocalizedModule p'.primeCompl + (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) + (LocalizedModule.mkLinearMap p'.primeCompl M))) + have : IsLocalizedModule (Algebra.algebraMapSubmonoid (Localization S) p'.primeCompl) l := + IsLocalizedModule.of_restrictScalars p'.primeCompl .. + apply IsLocalizedModule.of_exists_mul_mem + (Algebra.algebraMapSubmonoid (Localization S) p'.primeCompl) + · rintro _ ⟨x, hx, rfl⟩; exact hx + · rintro ⟨x, hx⟩ + obtain ⟨x, s, rfl⟩ := IsLocalization.mk'_surjective S x + refine ⟨algebraMap _ _ s.1, x, fun H ↦ hx ?_, by simp⟩ + rw [IsLocalization.mk'_eq_mul_mk'_one] + exact Ideal.mul_mem_right _ _ H + rw [mem_freeLocus_of_isLocalization (R := Localization S) p Rₚ Mₚ l] + rfl + +lemma freeLocus_eq_univ_iff [Module.FinitePresentation R M] : + freeLocus R M = Set.univ ↔ Module.Projective R M := by + simp_rw [Set.eq_univ_iff_forall, mem_freeLocus] + exact ⟨fun H ↦ Module.projective_of_localization_maximal fun I hI ↦ + have := H ⟨I, hI.isPrime⟩; .of_free, fun H x ↦ Module.free_of_flat_of_isLocalRing⟩ + +lemma freeLocus_eq_univ [Module.FinitePresentation R M] [Module.Flat R M] : + freeLocus R M = Set.univ := by + simp_rw [Set.eq_univ_iff_forall, mem_freeLocus] + exact fun x ↦ Module.free_of_flat_of_isLocalRing + +lemma basicOpen_subset_freeLocus_iff [Module.FinitePresentation R M] {f : R} : + (basicOpen f : Set (PrimeSpectrum R)) ⊆ freeLocus R M ↔ + Module.Projective (Localization.Away f) (LocalizedModule (.powers f) M) := by + rw [← freeLocus_eq_univ_iff, freeLocus_localization, + Set.preimage_eq_univ_iff, localization_away_comap_range _ f] + +lemma isOpen_freeLocus [Module.FinitePresentation R M] : + IsOpen (freeLocus R M) := by + refine isOpen_iff_forall_mem_open.mpr fun x hx ↦ ?_ + have : Module.Free _ _ := hx + obtain ⟨r, hr, hr', _⟩ := Module.FinitePresentation.exists_free_localizedModule_powers + x.asIdeal.primeCompl (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) + (Localization.AtPrime x.asIdeal) + exact ⟨basicOpen r, basicOpen_subset_freeLocus_iff.mpr inferInstance, (basicOpen r).2, hr⟩ + +variable (M) in +/-- The rank of `M` at the stalk of `p` is the rank of `Mₚ` as a `Rₚ`-module. -/ +noncomputable +def rankAtStalk (p : PrimeSpectrum R) : ℕ := + Module.finrank (Localization.AtPrime p.asIdeal) (LocalizedModule p.asIdeal.primeCompl M) + +lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : + IsLocallyConstant (fun x : freeLocus R M ↦ rankAtStalk M x.1) := by + refine (IsLocallyConstant.iff_exists_open _).mpr fun ⟨x, hx⟩ ↦ ?_ + have : Module.Free _ _ := hx + obtain ⟨f, hf, hf', hf''⟩ := Module.FinitePresentation.exists_free_localizedModule_powers + x.asIdeal.primeCompl (LocalizedModule.mkLinearMap x.asIdeal.primeCompl M) + (Localization.AtPrime x.asIdeal) + refine ⟨Subtype.val ⁻¹' basicOpen f, (basicOpen f).2.preimage continuous_subtype_val, hf, ?_⟩ + rintro ⟨p, hp''⟩ hp + let p' := Algebra.algebraMapSubmonoid (Localization (.powers f)) p.asIdeal.primeCompl + have hp' : Submonoid.powers f ≤ p.asIdeal.primeCompl := by + simpa [Submonoid.powers_le, Ideal.primeCompl] + let Rₚ := Localization.AtPrime p.asIdeal + let Mₚ := LocalizedModule p.asIdeal.primeCompl M + letI : Algebra (Localization.Away f) Rₚ := + IsLocalization.localizationAlgebraOfSubmonoidLe _ _ (.powers f) p.asIdeal.primeCompl hp' + have : IsScalarTower R (Localization.Away f) Rₚ := + IsLocalization.localization_isScalarTower_of_submonoid_le .. + letI : Module (Localization.Away f) Mₚ := Module.compHom Mₚ (algebraMap _ Rₚ) + have : IsScalarTower R (Localization.Away f) Mₚ := + ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by + simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ + have : IsScalarTower (Localization.Away f) Rₚ Mₚ := + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)).extendScalarsOfIsLocalization (.powers f) + (Localization.Away f) + have : IsLocalization p' Rₚ := + IsLocalization.isLocalization_of_submonoid_le (Localization.Away f) Rₚ _ _ hp' + have : IsLocalizedModule p.asIdeal.primeCompl (l.restrictScalars R) := + inferInstanceAs (IsLocalizedModule p.asIdeal.primeCompl + ((IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) + (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)))) + have : IsLocalizedModule (Algebra.algebraMapSubmonoid _ p.asIdeal.primeCompl) l := + IsLocalizedModule.of_restrictScalars p.asIdeal.primeCompl .. + have := Module.finrank_of_isLocalizedModule_of_free Rₚ p' l + simp [rankAtStalk, this, hf''] + +lemma isLocallyConstant_rankAtStalk [Module.FinitePresentation R M] [Module.Flat R M] : + IsLocallyConstant (rankAtStalk (R := R) M) := by + let e : freeLocus R M ≃ₜ PrimeSpectrum R := + (Homeomorph.setCongr freeLocus_eq_univ).trans (Homeomorph.Set.univ (PrimeSpectrum R)) + convert isLocallyConstant_rankAtStalk_freeLocus.comp_continuous e.symm.continuous + +end Module diff --git a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean index cd73b52dd3c65..efe62ee72caee 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Basic.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Basic.lean @@ -560,7 +560,7 @@ lemma IsLocalizedModule.eq_iff_exists [IsLocalizedModule S f] {x₁ x₂} : simp_rw [f.map_smul_of_tower, Submonoid.smul_def, ← Module.algebraMap_end_apply R R] at h exact ((Module.End_isUnit_iff _).mp <| map_units f c).1 h -theorem IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalizedModule S f] : +instance IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalizedModule S f] : IsLocalizedModule S (e ∘ₗ f : M →ₗ[R] M'') where map_units s := by rw [show algebraMap R (Module.End R M'') s = e ∘ₗ (algebraMap R (Module.End R M') s) ∘ₗ e.symm @@ -576,6 +576,18 @@ theorem IsLocalizedModule.of_linearEquiv (e : M' ≃ₗ[R] M'') [hf : IsLocalize EmbeddingLike.apply_eq_iff_eq] at h exact hf.exists_of_eq h +instance IsLocalizedModule.of_linearEquiv_right (e : M'' ≃ₗ[R] M) [hf : IsLocalizedModule S f] : + IsLocalizedModule S (f ∘ₗ e : M'' →ₗ[R] M') where + map_units s := hf.map_units s + surj' x := by + obtain ⟨⟨p, s⟩, h⟩ := hf.surj' x + exact ⟨⟨e.symm p, s⟩, by simpa using h⟩ + exists_of_eq h := by + simp_rw [LinearMap.coe_comp, LinearEquiv.coe_coe, Function.comp_apply] at h + obtain ⟨c, hc⟩ := hf.exists_of_eq h + exact ⟨c, by simpa only [Submonoid.smul_def, map_smul, e.symm_apply_apply] + using congr(e.symm $hc)⟩ + variable (M) in lemma isLocalizedModule_id (R') [CommSemiring R'] [Algebra R R'] [IsLocalization S R'] [Module R' M] [IsScalarTower R R' M] : IsLocalizedModule S (.id : M →ₗ[R] M) where @@ -703,6 +715,39 @@ instance localizedModuleIsLocalizedModule : LocalizedModule.mk_cancel t] exists_of_eq eq1 := by simpa only [eq_comm, one_smul] using LocalizedModule.mk_eq.mp eq1 +lemma IsLocalizedModule.of_restrictScalars (S : Submonoid R) + {N : Type*} [AddCommGroup N] [Module R N] [Module A M] [Module A N] + [IsScalarTower R A M] [IsScalarTower R A N] + (f : M →ₗ[A] N) [IsLocalizedModule S (f.restrictScalars R)] : + IsLocalizedModule (Algebra.algebraMapSubmonoid A S) f where + map_units x := by + obtain ⟨_, x, hx, rfl⟩ := x + have := IsLocalizedModule.map_units (f.restrictScalars R) ⟨x, hx⟩ + simp only [← IsScalarTower.algebraMap_apply, Module.End_isUnit_iff] at this ⊢ + exact this + surj' y := by + obtain ⟨⟨x, t⟩, e⟩ := IsLocalizedModule.surj S (f.restrictScalars R) y + exact ⟨⟨x, ⟨_, t, t.2, rfl⟩⟩, by simpa [Submonoid.smul_def] using e⟩ + exists_of_eq {x₁ x₂} e := by + obtain ⟨c, hc⟩ := IsLocalizedModule.exists_of_eq (S := S) (f := f.restrictScalars R) e + refine ⟨⟨_, c, c.2, rfl⟩, by simpa [Submonoid.smul_def]⟩ + +lemma IsLocalizedModule.of_exists_mul_mem {N : Type*} [AddCommGroup N] [Module R N] + (S T : Submonoid R) (h : S ≤ T) (h' : ∀ x : T, ∃ m : R, m * x ∈ S) + (f : M →ₗ[R] N) [IsLocalizedModule S f] : + IsLocalizedModule T f where + map_units x := by + obtain ⟨m, mx⟩ := h' x + have := IsLocalizedModule.map_units f ⟨_, mx⟩ + rw [map_mul, (Algebra.commute_algebraMap_left _ _).isUnit_mul_iff] at this + exact this.2 + surj' y := by + obtain ⟨⟨x, t⟩, e⟩ := IsLocalizedModule.surj S f y + exact ⟨⟨x, ⟨t, h t.2⟩⟩, e⟩ + exists_of_eq {x₁ x₂} e := by + obtain ⟨c, hc⟩ := IsLocalizedModule.exists_of_eq (S := S) (f := f) e + exact ⟨⟨c, h c.2⟩, hc⟩ + namespace IsLocalizedModule variable [IsLocalizedModule S f] diff --git a/Mathlib/LinearAlgebra/FreeModule/Basic.lean b/Mathlib/LinearAlgebra/FreeModule/Basic.lean index a53c69bdf4ff3..99e69913af94b 100644 --- a/Mathlib/LinearAlgebra/FreeModule/Basic.lean +++ b/Mathlib/LinearAlgebra/FreeModule/Basic.lean @@ -118,6 +118,27 @@ theorem of_equiv' {P : Type v} [AddCommMonoid P] [Module R P] (_ : Module.Free R (e : P ≃ₗ[R] N) : Module.Free R N := of_equiv e +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma of_ringEquiv {R R' M M'} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring R'] [AddCommMonoid M'] [Module R' M'] + (e₁ : R ≃+* R') (e₂ : M ≃ₛₗ[RingHomClass.toRingHom e₁] M') [Module.Free R M] : + Module.Free R' M' := by + let I := Module.Free.ChooseBasisIndex R M + obtain ⟨e₃ : M ≃ₗ[R] I →₀ R⟩ := Module.Free.chooseBasis R M + let e : M' ≃+ (I →₀ R') := + (e₂.symm.trans e₃).toAddEquiv.trans (Finsupp.mapRange.addEquiv (α := I) e₁.toAddEquiv) + have he (x) : e x = Finsupp.mapRange.addEquiv (α := I) e₁.toAddEquiv (e₃ (e₂.symm x)) := rfl + let e' : M' ≃ₗ[R'] (I →₀ R') := + { __ := e, map_smul' := fun m x ↦ Finsupp.ext fun i ↦ by simp [he, map_smulₛₗ] } + exact of_basis (.ofRepr e') + +attribute [local instance] RingHomInvPair.of_ringEquiv in +lemma iff_of_ringEquiv {R R' M M'} [Semiring R] [AddCommMonoid M] [Module R M] + [Semiring R'] [AddCommMonoid M'] [Module R' M'] + (e₁ : R ≃+* R') (e₂ : M ≃ₛₗ[RingHomClass.toRingHom e₁] M') : + Module.Free R M ↔ Module.Free R' M' := + ⟨fun _ ↦ of_ringEquiv e₁ e₂, fun _ ↦ of_ringEquiv e₁.symm e₂.symm⟩ + variable (R M N) /-- The module structure provided by `Semiring.toModule` is free. -/ diff --git a/Mathlib/RingTheory/Finiteness/Basic.lean b/Mathlib/RingTheory/Finiteness/Basic.lean index f68cde85ea80c..73248997ca909 100644 --- a/Mathlib/RingTheory/Finiteness/Basic.lean +++ b/Mathlib/RingTheory/Finiteness/Basic.lean @@ -119,6 +119,17 @@ theorem fg_restrictScalars {R S M : Type*} [CommSemiring R] [Semiring S] [Algebr use X exact (Submodule.restrictScalars_span R S h (X : Set M)).symm +lemma FG.of_restrictScalars (R) {A M} [CommSemiring R] [Semiring A] [AddCommMonoid M] + [Algebra R A] [Module R M] [Module A M] [IsScalarTower R A M] (S : Submodule A M) + (hS : (S.restrictScalars R).FG) : S.FG := by + obtain ⟨s, e⟩ := hS + refine ⟨s, Submodule.restrictScalars_injective R _ _ (le_antisymm ?_ ?_)⟩ + · show Submodule.span A s ≤ S + have := Submodule.span_le.mp e.le + rwa [Submodule.span_le] + · rw [← e] + exact Submodule.span_le_restrictScalars _ _ _ + theorem FG.stabilizes_of_iSup_eq {M' : Submodule R M} (hM' : M'.FG) (N : ℕ →o Submodule R M) (H : iSup N = M') : ∃ n, M' = N n := by obtain ⟨S, hS⟩ := hM' diff --git a/Mathlib/RingTheory/LocalProperties/Projective.lean b/Mathlib/RingTheory/LocalProperties/Projective.lean index a2c4615a7942e..d6c0268663ed7 100644 --- a/Mathlib/RingTheory/LocalProperties/Projective.lean +++ b/Mathlib/RingTheory/LocalProperties/Projective.lean @@ -5,9 +5,9 @@ Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.FinitePresentation import Mathlib.Algebra.Module.Projective -import Mathlib.LinearAlgebra.FreeModule.Basic +import Mathlib.LinearAlgebra.Dimension.Constructions +import Mathlib.LinearAlgebra.FreeModule.StrongRankCondition import Mathlib.RingTheory.LocalProperties.Submodule -import Mathlib.RingTheory.Localization.BaseChange /-! @@ -25,8 +25,43 @@ import Mathlib.RingTheory.Localization.BaseChange -/ -variable {R M N N'} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] -variable [AddCommGroup N'] [Module R N'] (S : Submonoid R) +universe uM + +variable {R N N' : Type*} {M : Type uM} [CommRing R] [AddCommGroup M] [Module R M] [AddCommGroup N] +variable [Module R N] [AddCommGroup N'] [Module R N'] (S : Submonoid R) + +theorem Module.free_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.Free R M] : + Module.Free Rₛ Mₛ := + Free.of_equiv (IsLocalizedModule.isBaseChange S Rₛ f).equiv + +universe uR' uM' in +/-- +Also see `IsLocalizedModule.lift_rank_eq` for a version for non-free modules, +but requires `S` to not contain any zero-divisors. +-/ +theorem Module.lift_rank_of_isLocalizedModule_of_free + (Rₛ : Type uR') {Mₛ : Type uM'} [AddCommGroup Mₛ] [Module R Mₛ] + [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S : Submonoid R) + (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Free R M] + [Nontrivial Rₛ] : + Cardinal.lift.{uM} (Module.rank Rₛ Mₛ) = Cardinal.lift.{uM'} (Module.rank R M) := by + apply Cardinal.lift_injective.{max uM' uR'} + have := (algebraMap R Rₛ).domain_nontrivial + have := (IsLocalizedModule.isBaseChange S Rₛ f).equiv.lift_rank_eq.symm + simp only [rank_tensorProduct, rank_self, + Cardinal.lift_one, one_mul, Cardinal.lift_lift] at this ⊢ + convert this + exact Cardinal.lift_umax + +theorem Module.finrank_of_isLocalizedModule_of_free + (Rₛ : Type*) {Mₛ : Type*} [AddCommGroup Mₛ] [Module R Mₛ] + [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] (S : Submonoid R) + (f : M →ₗ[R] Mₛ) [IsLocalization S Rₛ] [IsLocalizedModule S f] [Module.Free R M] + [Nontrivial Rₛ] : + Module.finrank Rₛ Mₛ = Module.finrank R M := by + simpa using congr(Cardinal.toNat $(Module.lift_rank_of_isLocalizedModule_of_free Rₛ S f)) theorem Module.projective_of_isLocalizedModule {Rₛ Mₛ} [AddCommGroup Mₛ] [Module R Mₛ] [CommRing Rₛ] [Algebra R Rₛ] [Module Rₛ Mₛ] [IsScalarTower R Rₛ Mₛ] From 916ff700f9f452a36f3b6f37ace0f367c5f680f8 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 19:46:35 +1100 Subject: [PATCH 120/250] bump and fix --- .../Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean | 4 ++-- lake-manifest.json | 2 +- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean index 17812e9c56178..4ecfb36c3cab1 100644 --- a/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean +++ b/Mathlib/Tactic/Linarith/Oracle/SimplexAlgorithm/SimplexAlgorithm.lean @@ -48,8 +48,8 @@ def doPivotOperation (exitIdx enterIdx : Nat) : SimplexAlgorithmM matType Unit : let newBasic := s.basic.set! exitIdx s.free[enterIdx]! let newFree := s.free.set! enterIdx s.basic[exitIdx]! - have hb : newBasic.size = s.basic.size := by apply Array.size_setD - have hf : newFree.size = s.free.size := by apply Array.size_setD + have hb : newBasic.size = s.basic.size := by apply Array.size_setIfInBounds + have hf : newFree.size = s.free.size := by apply Array.size_setIfInBounds return (⟨newBasic, newFree, hb ▸ hf ▸ mat⟩ : Tableau matType) diff --git a/lake-manifest.json b/lake-manifest.json index 0e0184ac9c4d6..cdc932a7c6e5d 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "dc439388cc078925934b1c3e06276194280822d8", + "rev": "0681948d281fb8d3499952cd954be673c17e6e81", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From 454bc8a132d5ad572487513d801e41fc162dc128 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 09:04:44 +0000 Subject: [PATCH 121/250] chore: split Order.Filter.Basic, creating Order.Filter.Finite (#19354) --- Mathlib.lean | 1 + Mathlib/Order/Filter/Bases.lean | 2 +- Mathlib/Order/Filter/Basic.lean | 311 +-------------- Mathlib/Order/Filter/CardinalInter.lean | 1 + Mathlib/Order/Filter/Cocardinal.lean | 1 - .../Order/Filter/CountableSeparatingOn.lean | 2 +- Mathlib/Order/Filter/Extr.lean | 1 + Mathlib/Order/Filter/Finite.lean | 359 ++++++++++++++++++ Mathlib/Order/Filter/Germ/Basic.lean | 1 + Mathlib/Order/Filter/Subsingleton.lean | 1 - .../SetTheory/Cardinal/CountableCover.lean | 2 +- 11 files changed, 380 insertions(+), 302 deletions(-) create mode 100644 Mathlib/Order/Filter/Finite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 7e32e8e68409c..56e4a03916d8e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3953,6 +3953,7 @@ import Mathlib.Order.Filter.ENNReal import Mathlib.Order.Filter.EventuallyConst import Mathlib.Order.Filter.Extr import Mathlib.Order.Filter.FilterProduct +import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.Germ.Basic import Mathlib.Order.Filter.Germ.OrderedMonoid import Mathlib.Order.Filter.IndicatorFunction diff --git a/Mathlib/Order/Filter/Bases.lean b/Mathlib/Order/Filter/Bases.lean index ffd5fe2a6c9e4..ab8dd13872eb6 100644 --- a/Mathlib/Order/Filter/Bases.lean +++ b/Mathlib/Order/Filter/Bases.lean @@ -5,7 +5,7 @@ Authors: Yury Kudryashov, Johannes Hölzl, Mario Carneiro, Patrick Massot -/ import Mathlib.Data.Prod.PProd import Mathlib.Data.Set.Countable -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Finite /-! # Filter bases diff --git a/Mathlib/Order/Filter/Basic.lean b/Mathlib/Order/Filter/Basic.lean index 995ddfea3234b..65c9595abbf40 100644 --- a/Mathlib/Order/Filter/Basic.lean +++ b/Mathlib/Order/Filter/Basic.lean @@ -3,7 +3,10 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Jeremy Avigad -/ -import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Algebra.Group.Basic +import Mathlib.Algebra.Group.Pi.Basic +import Mathlib.Control.Basic +import Mathlib.Data.Set.Lattice import Mathlib.Order.Filter.Defs /-! @@ -59,6 +62,7 @@ we do *not* require. This gives `Filter X` better formal properties, in particul -/ assert_not_exists OrderedSemiring +assert_not_exists Fintype open Function Set Order open scoped symmDiff @@ -103,27 +107,15 @@ theorem congr_sets (h : { x | x ∈ s ↔ x ∈ t } ∈ f) : s ∈ f ↔ t ∈ f lemma copy_eq {S} (hmem : ∀ s, s ∈ S ↔ s ∈ f) : f.copy S hmem = f := Filter.ext hmem -@[simp] -theorem biInter_mem {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Finite) : - (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := - Finite.induction_on hf (by simp) fun _ _ hs => by simp [hs] - -@[simp] -theorem biInter_finset_mem {β : Type v} {s : β → Set α} (is : Finset β) : - (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := - biInter_mem is.finite_toSet - -alias _root_.Finset.iInter_mem_sets := biInter_finset_mem +/-- Weaker version of `Filter.biInter_mem` that assumes `Subsingleton β` rather than `Finite β`. -/ +theorem biInter_mem' {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Subsingleton) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := by + apply Subsingleton.induction_on hf <;> simp --- attribute [protected] Finset.iInter_mem_sets porting note: doesn't work - -@[simp] -theorem sInter_mem {s : Set (Set α)} (hfin : s.Finite) : ⋂₀ s ∈ f ↔ ∀ U ∈ s, U ∈ f := by - rw [sInter_eq_biInter, biInter_mem hfin] - -@[simp] -theorem iInter_mem {β : Sort v} {s : β → Set α} [Finite β] : (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := - (sInter_mem (finite_range _)).trans forall_mem_range +/-- Weaker version of `Filter.iInter_mem` that assumes `Subsingleton β` rather than `Finite β`. -/ +theorem iInter_mem' {β : Sort v} {s : β → Set α} [Subsingleton β] : + (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := by + rw [← sInter_range, sInter_eq_biInter, biInter_mem' (subsingleton_range s), forall_mem_range] theorem exists_mem_subset_iff : (∃ t ∈ f, t ⊆ s) ↔ s ∈ f := ⟨fun ⟨_, ht, ts⟩ => mem_of_superset ht ts, fun hs => ⟨s, hs, Subset.rfl⟩⟩ @@ -181,24 +173,6 @@ theorem le_generate_iff {s : Set (Set α)} {f : Filter α} : f ≤ generate s hu.recOn (fun h' => h h') univ_mem (fun _ hxy hx => mem_of_superset hx hxy) fun _ _ hx hy => inter_mem hx hy -theorem mem_generate_iff {s : Set <| Set α} {U : Set α} : - U ∈ generate s ↔ ∃ t ⊆ s, Set.Finite t ∧ ⋂₀ t ⊆ U := by - constructor <;> intro h - · induction h with - | @basic V V_in => - exact ⟨{V}, singleton_subset_iff.2 V_in, finite_singleton _, (sInter_singleton _).subset⟩ - | univ => exact ⟨∅, empty_subset _, finite_empty, subset_univ _⟩ - | superset _ hVW hV => - rcases hV with ⟨t, hts, ht, htV⟩ - exact ⟨t, hts, ht, htV.trans hVW⟩ - | inter _ _ hV hW => - rcases hV, hW with ⟨⟨t, hts, ht, htV⟩, u, hus, hu, huW⟩ - exact - ⟨t ∪ u, union_subset hts hus, ht.union hu, - (sInter_union _ _).subset.trans <| inter_subset_inter htV huW⟩ - · rcases h with ⟨t, hts, tfin, h⟩ - exact mem_of_superset ((sInter_mem tfin).2 fun V hV => GenerateSets.basic <| hts hV) h - @[simp] lemma generate_singleton (s : Set α) : generate {s} = 𝓟 s := le_antisymm (fun _t ht ↦ mem_of_superset (mem_generate_of_mem <| mem_singleton _) ht) <| le_generate_iff.2 <| singleton_subset_iff.2 Subset.rfl @@ -334,58 +308,6 @@ theorem iInf_eq_generate (s : ι → Filter α) : iInf s = generate (⋃ i, (s i theorem mem_iInf_of_mem {f : ι → Filter α} (i : ι) {s} (hs : s ∈ f i) : s ∈ ⨅ i, f i := iInf_le f i hs -theorem mem_iInf_of_iInter {ι} {s : ι → Filter α} {U : Set α} {I : Set ι} (I_fin : I.Finite) - {V : I → Set α} (hV : ∀ (i : I), V i ∈ s i) (hU : ⋂ i, V i ⊆ U) : U ∈ ⨅ i, s i := by - haveI := I_fin.fintype - refine mem_of_superset (iInter_mem.2 fun i => ?_) hU - exact mem_iInf_of_mem (i : ι) (hV _) - -theorem mem_iInf {ι} {s : ι → Filter α} {U : Set α} : - (U ∈ ⨅ i, s i) ↔ - ∃ I : Set ι, I.Finite ∧ ∃ V : I → Set α, (∀ (i : I), V i ∈ s i) ∧ U = ⋂ i, V i := by - constructor - · rw [iInf_eq_generate, mem_generate_iff] - rintro ⟨t, tsub, tfin, tinter⟩ - rcases eq_finite_iUnion_of_finite_subset_iUnion tfin tsub with ⟨I, Ifin, σ, σfin, σsub, rfl⟩ - rw [sInter_iUnion] at tinter - set V := fun i => U ∪ ⋂₀ σ i with hV - have V_in : ∀ (i : I), V i ∈ s i := by - rintro i - have : ⋂₀ σ i ∈ s i := by - rw [sInter_mem (σfin _)] - apply σsub - exact mem_of_superset this subset_union_right - refine ⟨I, Ifin, V, V_in, ?_⟩ - rwa [hV, ← union_iInter, union_eq_self_of_subset_right] - · rintro ⟨I, Ifin, V, V_in, rfl⟩ - exact mem_iInf_of_iInter Ifin V_in Subset.rfl - -theorem mem_iInf' {ι} {s : ι → Filter α} {U : Set α} : - (U ∈ ⨅ i, s i) ↔ - ∃ I : Set ι, I.Finite ∧ ∃ V : ι → Set α, (∀ i, V i ∈ s i) ∧ - (∀ i ∉ I, V i = univ) ∧ (U = ⋂ i ∈ I, V i) ∧ U = ⋂ i, V i := by - classical - simp only [mem_iInf, SetCoe.forall', biInter_eq_iInter] - refine ⟨?_, fun ⟨I, If, V, hVs, _, hVU, _⟩ => ⟨I, If, fun i => V i, fun i => hVs i, hVU⟩⟩ - rintro ⟨I, If, V, hV, rfl⟩ - refine ⟨I, If, fun i => if hi : i ∈ I then V ⟨i, hi⟩ else univ, fun i => ?_, fun i hi => ?_, ?_⟩ - · dsimp only - split_ifs - exacts [hV ⟨i,_⟩, univ_mem] - · exact dif_neg hi - · simp only [iInter_dite, biInter_eq_iInter, dif_pos (Subtype.coe_prop _), Subtype.coe_eta, - iInter_univ, inter_univ, eq_self_iff_true, true_and] - -theorem exists_iInter_of_mem_iInf {ι : Type*} {α : Type*} {f : ι → Filter α} {s} - (hs : s ∈ ⨅ i, f i) : ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := - let ⟨_, _, V, hVs, _, _, hVU'⟩ := mem_iInf'.1 hs; ⟨V, hVs, hVU'⟩ - -theorem mem_iInf_of_finite {ι : Type*} [Finite ι] {α : Type*} {f : ι → Filter α} (s) : - (s ∈ ⨅ i, f i) ↔ ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := by - refine ⟨exists_iInter_of_mem_iInf, ?_⟩ - rintro ⟨t, ht, rfl⟩ - exact iInter_mem.2 fun i => mem_iInf_of_mem i (ht i) - @[simp] theorem le_principal_iff {s : Set α} {f : Filter α} : f ≤ 𝓟 s ↔ s ∈ f := ⟨fun h => h Subset.rfl, fun hs _ ht => mem_of_superset hs ht⟩ @@ -453,26 +375,6 @@ theorem NeBot.not_disjoint (hf : f.NeBot) (hs : s ∈ f) (ht : t ∈ f) : ¬Disj theorem inf_eq_bot_iff {f g : Filter α} : f ⊓ g = ⊥ ↔ ∃ U ∈ f, ∃ V ∈ g, U ∩ V = ∅ := by simp only [← disjoint_iff, Filter.disjoint_iff, Set.disjoint_iff_inter_eq_empty] -theorem _root_.Pairwise.exists_mem_filter_of_disjoint {ι : Type*} [Finite ι] {l : ι → Filter α} - (hd : Pairwise (Disjoint on l)) : - ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ Pairwise (Disjoint on s) := by - have : Pairwise fun i j => ∃ (s : {s // s ∈ l i}) (t : {t // t ∈ l j}), Disjoint s.1 t.1 := by - simpa only [Pairwise, Function.onFun, Filter.disjoint_iff, exists_prop, Subtype.exists] using hd - choose! s t hst using this - refine ⟨fun i => ⋂ j, @s i j ∩ @t j i, fun i => ?_, fun i j hij => ?_⟩ - exacts [iInter_mem.2 fun j => inter_mem (@s i j).2 (@t j i).2, - (hst hij).mono ((iInter_subset _ j).trans inter_subset_left) - ((iInter_subset _ i).trans inter_subset_right)] - -theorem _root_.Set.PairwiseDisjoint.exists_mem_filter {ι : Type*} {l : ι → Filter α} {t : Set ι} - (hd : t.PairwiseDisjoint l) (ht : t.Finite) : - ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ t.PairwiseDisjoint s := by - haveI := ht.to_subtype - rcases (hd.subtype _ _).exists_mem_filter_of_disjoint with ⟨s, hsl, hsd⟩ - lift s to (i : t) → {s // s ∈ l i} using hsl - rcases @Subtype.exists_pi_extension ι (fun i => { s // s ∈ l i }) _ _ s with ⟨s, rfl⟩ - exact ⟨fun i => s i, fun i => (s i).2, hsd.set_of_subtype _ _⟩ - /-- There is exactly one filter on an empty type. -/ instance unique [IsEmpty α] : Unique (Filter α) where default := ⊥ @@ -548,23 +450,6 @@ theorem biInf_sets_eq {f : β → Filter α} {s : Set β} (h : DirectedOn (f ⁻ (ne : s.Nonempty) : (⨅ i ∈ s, f i).sets = ⋃ i ∈ s, (f i).sets := ext fun t => by simp [mem_biInf_of_directed h ne] -theorem iInf_sets_eq_finite {ι : Type*} (f : ι → Filter α) : - (⨅ i, f i).sets = ⋃ t : Finset ι, (⨅ i ∈ t, f i).sets := by - rw [iInf_eq_iInf_finset, iInf_sets_eq] - exact directed_of_isDirected_le fun _ _ => biInf_mono - -theorem iInf_sets_eq_finite' (f : ι → Filter α) : - (⨅ i, f i).sets = ⋃ t : Finset (PLift ι), (⨅ i ∈ t, f (PLift.down i)).sets := by - rw [← iInf_sets_eq_finite, ← Equiv.plift.surjective.iInf_comp, Equiv.plift_apply] - -theorem mem_iInf_finite {ι : Type*} {f : ι → Filter α} (s) : - s ∈ iInf f ↔ ∃ t : Finset ι, s ∈ ⨅ i ∈ t, f i := - (Set.ext_iff.1 (iInf_sets_eq_finite f) s).trans mem_iUnion - -theorem mem_iInf_finite' {f : ι → Filter α} (s) : - s ∈ iInf f ↔ ∃ t : Finset (PLift ι), s ∈ ⨅ i ∈ t, f (PLift.down i) := - (Set.ext_iff.1 (iInf_sets_eq_finite' f) s).trans mem_iUnion - @[simp] theorem sup_join {f₁ f₂ : Filter (Filter α)} : join f₁ ⊔ join f₂ = join (f₁ ⊔ f₂) := Filter.ext fun x => by simp only [mem_sup, mem_join] @@ -583,38 +468,6 @@ instance : DistribLattice (Filter α) := ⟨t₁, x.sets_of_superset hs inter_subset_left, ht₁, t₂, x.sets_of_superset hs inter_subset_right, ht₂, rfl⟩ } -/-- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ --- See note [reducible non-instances] -abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := - { Filter.instCompleteLatticeFilter with - iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by - classical - rw [iInf_subtype'] - rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ - obtain ⟨u, hu⟩ := h₂ - rw [← Finset.inf_eq_iInf] at hu - suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ - refine Finset.induction_on u (le_sup_of_le_right le_top) ?_ - rintro ⟨i⟩ u _ ih - rw [Finset.inf_insert, sup_inf_left] - exact le_inf (iInf_le _ _) ih } - -instance instCoframe : Coframe (Filter α) := .ofMinimalAxioms coframeMinimalAxioms - -theorem mem_iInf_finset {s : Finset α} {f : α → Filter β} {t : Set β} : - (t ∈ ⨅ a ∈ s, f a) ↔ ∃ p : α → Set β, (∀ a ∈ s, p a ∈ f a) ∧ t = ⋂ a ∈ s, p a := by - classical - simp only [← Finset.set_biInter_coe, biInter_eq_iInter, iInf_subtype'] - refine ⟨fun h => ?_, ?_⟩ - · rcases (mem_iInf_of_finite _).1 h with ⟨p, hp, rfl⟩ - refine ⟨fun a => if h : a ∈ s then p ⟨a, h⟩ else univ, - fun a ha => by simpa [ha] using hp ⟨a, ha⟩, ?_⟩ - refine iInter_congr_of_surjective id surjective_id ?_ - rintro ⟨a, ha⟩ - simp [ha] - · rintro ⟨p, hpf, rfl⟩ - exact iInter_mem.2 fun a => mem_iInf_of_mem a (hpf a a.2) - /-- If `f : ι → Filter α` is directed, `ι` is not empty, and `∀ i, f i ≠ ⊥`, then `iInf f ≠ ⊥`. See also `iInf_neBot_of_directed` for a version assuming `Nonempty α` instead of `Nonempty ι`. -/ theorem iInf_neBot_of_directed' {f : ι → Filter α} [Nonempty ι] (hd : Directed (· ≥ ·) f) : @@ -650,20 +503,6 @@ theorem iInf_neBot_iff_of_directed {f : ι → Filter α} [Nonempty α] (hd : Di NeBot (iInf f) ↔ ∀ i, NeBot (f i) := ⟨fun H i => H.mono (iInf_le _ i), iInf_neBot_of_directed hd⟩ -@[elab_as_elim] -theorem iInf_sets_induct {f : ι → Filter α} {s : Set α} (hs : s ∈ iInf f) {p : Set α → Prop} - (uni : p univ) (ins : ∀ {i s₁ s₂}, s₁ ∈ f i → p s₂ → p (s₁ ∩ s₂)) : p s := by - classical - rw [mem_iInf_finite'] at hs - simp only [← Finset.inf_eq_iInf] at hs - rcases hs with ⟨is, his⟩ - induction is using Finset.induction_on generalizing s with - | empty => rwa [mem_top.1 his] - | insert _ ih => - rw [Finset.inf_insert, mem_inf_iff] at his - rcases his with ⟨s₁, hs₁, s₂, hs₂, rfl⟩ - exact ins hs₁ (ih hs₂) - /-! #### `principal` equations -/ @[simp] @@ -719,29 +558,6 @@ theorem diff_mem_inf_principal_compl {f : Filter α} {s : Set α} (hs : s ∈ f) theorem principal_le_iff {s : Set α} {f : Filter α} : 𝓟 s ≤ f ↔ ∀ V ∈ f, s ⊆ V := by simp_rw [le_def, mem_principal] -@[simp] -theorem iInf_principal_finset {ι : Type w} (s : Finset ι) (f : ι → Set α) : - ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by - classical - induction' s using Finset.induction_on with i s _ hs - · simp - · rw [Finset.iInf_insert, Finset.set_biInter_insert, hs, inf_principal] - -theorem iInf_principal {ι : Sort w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := by - cases nonempty_fintype (PLift ι) - rw [← iInf_plift_down, ← iInter_plift_down] - simpa using iInf_principal_finset Finset.univ (f <| PLift.down ·) - -/-- A special case of `iInf_principal` that is safe to mark `simp`. -/ -@[simp] -theorem iInf_principal' {ι : Type w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := - iInf_principal _ - -theorem iInf_principal_finite {ι : Type w} {s : Set ι} (hs : s.Finite) (f : ι → Set α) : - ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by - lift s to Finset ι using hs - exact mod_cast iInf_principal_finset s f - end Lattice @[mono, gcongr] @@ -820,28 +636,6 @@ theorem eventually_congr {f : Filter α} {p q : α → Prop} (h : ∀ᶠ x in f, (∀ᶠ x in f, p x) ↔ ∀ᶠ x in f, q x := ⟨fun hp => hp.congr h, fun hq => hq.congr <| by simpa only [Iff.comm] using h⟩ -@[simp] -theorem eventually_all {ι : Sort*} [Finite ι] {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i, p i x) ↔ ∀ i, ∀ᶠ x in l, p i x := by - simpa only [Filter.Eventually, setOf_forall] using iInter_mem - -@[simp] -theorem eventually_all_finite {ι} {I : Set ι} (hI : I.Finite) {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := by - simpa only [Filter.Eventually, setOf_forall] using biInter_mem hI - -alias _root_.Set.Finite.eventually_all := eventually_all_finite - --- attribute [protected] Set.Finite.eventually_all - -@[simp] theorem eventually_all_finset {ι} (I : Finset ι) {l} {p : ι → α → Prop} : - (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := - I.finite_toSet.eventually_all - -alias _root_.Finset.eventually_all := eventually_all_finset - --- attribute [protected] Finset.eventually_all - @[simp] theorem eventually_or_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : (∀ᶠ x in f, p ∨ q x) ↔ p ∨ ∀ᶠ x in f, q x := @@ -852,10 +646,6 @@ theorem eventually_or_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} (∀ᶠ x in f, p x ∨ q) ↔ (∀ᶠ x in f, p x) ∨ q := by simp only [@or_comm _ q, eventually_or_distrib_left] -theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := - eventually_all - @[simp] theorem eventually_bot {p : α → Prop} : ∀ᶠ x in ⊥, p x := ⟨⟩ @@ -1008,16 +798,6 @@ theorem eventually_imp_distrib_right {f : Filter α} {p : α → Prop} {q : Prop (∀ᶠ x in f, p x → q) ↔ (∃ᶠ x in f, p x) → q := by simp only [imp_iff_not_or, eventually_or_distrib_right, not_frequently] -@[simp] -theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : - (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by - simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] - -@[simp] -theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : - (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by - simp only [@and_comm _ q, frequently_and_distrib_left] - @[simp] theorem frequently_bot {p : α → Prop} : ¬∃ᶠ x in ⊥, p x := by simp @@ -1328,69 +1108,6 @@ theorem EventuallyLE.union {s t s' t' : Set α} {l : Filter α} (h : s ≤ᶠ[l] (s ∪ s' : Set α) ≤ᶠ[l] (t ∪ t' : Set α) := h'.mp <| h.mono fun _ => Or.imp -protected lemma EventuallyLE.iUnion [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := - (eventually_all.2 h).mono fun _x hx hx' ↦ - let ⟨i, hi⟩ := mem_iUnion.1 hx'; mem_iUnion.2 ⟨i, hx i hi⟩ - -protected lemma EventuallyEq.iUnion [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := - (EventuallyLE.iUnion fun i ↦ (h i).le).antisymm <| .iUnion fun i ↦ (h i).symm.le - -protected lemma EventuallyLE.iInter [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := - (eventually_all.2 h).mono fun _x hx hx' ↦ mem_iInter.2 fun i ↦ hx i (mem_iInter.1 hx' i) - -protected lemma EventuallyEq.iInter [Finite ι] {s t : ι → Set α} - (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := - (EventuallyLE.iInter fun i ↦ (h i).le).antisymm <| .iInter fun i ↦ (h i).symm.le - -lemma _root_.Set.Finite.eventuallyLE_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := by - have := hs.to_subtype - rw [biUnion_eq_iUnion, biUnion_eq_iUnion] - exact .iUnion fun i ↦ hle i.1 i.2 - -alias EventuallyLE.biUnion := Set.Finite.eventuallyLE_iUnion - -lemma _root_.Set.Finite.eventuallyEq_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := - (EventuallyLE.biUnion hs fun i hi ↦ (heq i hi).le).antisymm <| - .biUnion hs fun i hi ↦ (heq i hi).symm.le - -alias EventuallyEq.biUnion := Set.Finite.eventuallyEq_iUnion - -lemma _root_.Set.Finite.eventuallyLE_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := by - have := hs.to_subtype - rw [biInter_eq_iInter, biInter_eq_iInter] - exact .iInter fun i ↦ hle i.1 i.2 - -alias EventuallyLE.biInter := Set.Finite.eventuallyLE_iInter - -lemma _root_.Set.Finite.eventuallyEq_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) - {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := - (EventuallyLE.biInter hs fun i hi ↦ (heq i hi).le).antisymm <| - .biInter hs fun i hi ↦ (heq i hi).symm.le - -alias EventuallyEq.biInter := Set.Finite.eventuallyEq_iInter - -lemma _root_.Finset.eventuallyLE_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := - .biUnion s.finite_toSet hle - -lemma _root_.Finset.eventuallyEq_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := - .biUnion s.finite_toSet heq - -lemma _root_.Finset.eventuallyLE_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := - .biInter s.finite_toSet hle - -lemma _root_.Finset.eventuallyEq_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} - (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := - .biInter s.finite_toSet heq - @[mono] theorem EventuallyLE.compl {s t : Set α} {l : Filter α} (h : s ≤ᶠ[l] t) : (tᶜ : Set α) ≤ᶠ[l] (sᶜ : Set α) := @@ -2372,4 +2089,4 @@ lemma compl_mem_comk {p : Set α → Prop} {he hmono hunion s} : end Filter -set_option linter.style.longFile 2500 +set_option linter.style.longFile 2200 diff --git a/Mathlib/Order/Filter/CardinalInter.lean b/Mathlib/Order/Filter/CardinalInter.lean index 22967d8a816b3..9a1e5bb744169 100644 --- a/Mathlib/Order/Filter/CardinalInter.lean +++ b/Mathlib/Order/Filter/CardinalInter.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Josha Dekker -/ import Mathlib.Order.Filter.Tendsto +import Mathlib.Order.Filter.Finite import Mathlib.Order.Filter.CountableInter import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality diff --git a/Mathlib/Order/Filter/Cocardinal.lean b/Mathlib/Order/Filter/Cocardinal.lean index 25e10e20a8fa3..d3da9a3ef6505 100644 --- a/Mathlib/Order/Filter/Cocardinal.lean +++ b/Mathlib/Order/Filter/Cocardinal.lean @@ -8,7 +8,6 @@ import Mathlib.Order.Filter.CountableInter import Mathlib.Order.Filter.CardinalInter import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Cofinality -import Mathlib.Order.Filter.Bases /-! # The cocardinal filter diff --git a/Mathlib/Order/Filter/CountableSeparatingOn.lean b/Mathlib/Order/Filter/CountableSeparatingOn.lean index b1f268cda8419..552a128f17bc5 100644 --- a/Mathlib/Order/Filter/CountableSeparatingOn.lean +++ b/Mathlib/Order/Filter/CountableSeparatingOn.lean @@ -167,7 +167,7 @@ theorem exists_subset_subsingleton_mem_of_forall_separating (p : Set α → Prop | inr hsl => simp only [hx.2 s hsS hsl, hy.2 s hsS hsl] · exact inter_mem (inter_mem hs ((countable_sInter_mem (hSc.mono inter_subset_left)).2 fun _ h ↦ h.2)) - ((countable_bInter_mem hSc).2 fun U hU ↦ iInter_mem.2 id) + ((countable_bInter_mem hSc).2 fun U hU ↦ iInter_mem'.2 id) theorem exists_mem_singleton_mem_of_mem_of_nonempty_of_forall_separating (p : Set α → Prop) {s : Set α} [HasCountableSeparatingOn α p s] (hs : s ∈ l) (hne : s.Nonempty) diff --git a/Mathlib/Order/Filter/Extr.lean b/Mathlib/Order/Filter/Extr.lean index ba1e7c93d0896..5f405e77557b1 100644 --- a/Mathlib/Order/Filter/Extr.lean +++ b/Mathlib/Order/Filter/Extr.lean @@ -6,6 +6,7 @@ Authors: Yury Kudryashov import Mathlib.Order.Filter.Tendsto import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Algebra.Order.Group.Defs +import Mathlib.Data.Finset.Lattice.Fold /-! # Minimum and maximum w.r.t. a filter and on a set diff --git a/Mathlib/Order/Filter/Finite.lean b/Mathlib/Order/Filter/Finite.lean new file mode 100644 index 0000000000000..2c96231b9a89f --- /dev/null +++ b/Mathlib/Order/Filter/Finite.lean @@ -0,0 +1,359 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl, Jeremy Avigad +-/ +import Mathlib.Data.Set.Finite.Lattice +import Mathlib.Order.Filter.Basic + +/-! +# Results filters related to finiteness. + +-/ + + + +open Function Set Order +open scoped symmDiff + +universe u v w x y + +namespace Filter + +variable {α : Type u} {f g : Filter α} {s t : Set α} + +@[simp] +theorem biInter_mem {β : Type v} {s : β → Set α} {is : Set β} (hf : is.Finite) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := + Finite.induction_on hf (by simp) fun _ _ hs => by simp [hs] + +@[simp] +theorem biInter_finset_mem {β : Type v} {s : β → Set α} (is : Finset β) : + (⋂ i ∈ is, s i) ∈ f ↔ ∀ i ∈ is, s i ∈ f := + biInter_mem is.finite_toSet + +alias _root_.Finset.iInter_mem_sets := biInter_finset_mem + +-- attribute [protected] Finset.iInter_mem_sets porting note: doesn't work + +@[simp] +theorem sInter_mem {s : Set (Set α)} (hfin : s.Finite) : ⋂₀ s ∈ f ↔ ∀ U ∈ s, U ∈ f := by + rw [sInter_eq_biInter, biInter_mem hfin] + +@[simp] +theorem iInter_mem {β : Sort v} {s : β → Set α} [Finite β] : (⋂ i, s i) ∈ f ↔ ∀ i, s i ∈ f := + (sInter_mem (finite_range _)).trans forall_mem_range + +end Filter + + +namespace Filter + +variable {α : Type u} {β : Type v} {γ : Type w} {δ : Type*} {ι : Sort x} + +section Lattice + +variable {f g : Filter α} {s t : Set α} + +theorem mem_generate_iff {s : Set <| Set α} {U : Set α} : + U ∈ generate s ↔ ∃ t ⊆ s, Set.Finite t ∧ ⋂₀ t ⊆ U := by + constructor <;> intro h + · induction h with + | @basic V V_in => + exact ⟨{V}, singleton_subset_iff.2 V_in, finite_singleton _, (sInter_singleton _).subset⟩ + | univ => exact ⟨∅, empty_subset _, finite_empty, subset_univ _⟩ + | superset _ hVW hV => + rcases hV with ⟨t, hts, ht, htV⟩ + exact ⟨t, hts, ht, htV.trans hVW⟩ + | inter _ _ hV hW => + rcases hV, hW with ⟨⟨t, hts, ht, htV⟩, u, hus, hu, huW⟩ + exact + ⟨t ∪ u, union_subset hts hus, ht.union hu, + (sInter_union _ _).subset.trans <| inter_subset_inter htV huW⟩ + · rcases h with ⟨t, hts, tfin, h⟩ + exact mem_of_superset ((sInter_mem tfin).2 fun V hV => GenerateSets.basic <| hts hV) h + +theorem mem_iInf_of_iInter {ι} {s : ι → Filter α} {U : Set α} {I : Set ι} (I_fin : I.Finite) + {V : I → Set α} (hV : ∀ (i : I), V i ∈ s i) (hU : ⋂ i, V i ⊆ U) : U ∈ ⨅ i, s i := by + haveI := I_fin.fintype + refine mem_of_superset (iInter_mem.2 fun i => ?_) hU + exact mem_iInf_of_mem (i : ι) (hV _) + +theorem mem_iInf {ι} {s : ι → Filter α} {U : Set α} : + (U ∈ ⨅ i, s i) ↔ + ∃ I : Set ι, I.Finite ∧ ∃ V : I → Set α, (∀ (i : I), V i ∈ s i) ∧ U = ⋂ i, V i := by + constructor + · rw [iInf_eq_generate, mem_generate_iff] + rintro ⟨t, tsub, tfin, tinter⟩ + rcases eq_finite_iUnion_of_finite_subset_iUnion tfin tsub with ⟨I, Ifin, σ, σfin, σsub, rfl⟩ + rw [sInter_iUnion] at tinter + set V := fun i => U ∪ ⋂₀ σ i with hV + have V_in : ∀ (i : I), V i ∈ s i := by + rintro i + have : ⋂₀ σ i ∈ s i := by + rw [sInter_mem (σfin _)] + apply σsub + exact mem_of_superset this subset_union_right + refine ⟨I, Ifin, V, V_in, ?_⟩ + rwa [hV, ← union_iInter, union_eq_self_of_subset_right] + · rintro ⟨I, Ifin, V, V_in, rfl⟩ + exact mem_iInf_of_iInter Ifin V_in Subset.rfl + +theorem mem_iInf' {ι} {s : ι → Filter α} {U : Set α} : + (U ∈ ⨅ i, s i) ↔ + ∃ I : Set ι, I.Finite ∧ ∃ V : ι → Set α, (∀ i, V i ∈ s i) ∧ + (∀ i ∉ I, V i = univ) ∧ (U = ⋂ i ∈ I, V i) ∧ U = ⋂ i, V i := by + classical + simp only [mem_iInf, SetCoe.forall', biInter_eq_iInter] + refine ⟨?_, fun ⟨I, If, V, hVs, _, hVU, _⟩ => ⟨I, If, fun i => V i, fun i => hVs i, hVU⟩⟩ + rintro ⟨I, If, V, hV, rfl⟩ + refine ⟨I, If, fun i => if hi : i ∈ I then V ⟨i, hi⟩ else univ, fun i => ?_, fun i hi => ?_, ?_⟩ + · dsimp only + split_ifs + exacts [hV ⟨i,_⟩, univ_mem] + · exact dif_neg hi + · simp only [iInter_dite, biInter_eq_iInter, dif_pos (Subtype.coe_prop _), Subtype.coe_eta, + iInter_univ, inter_univ, eq_self_iff_true, true_and] + +theorem exists_iInter_of_mem_iInf {ι : Type*} {α : Type*} {f : ι → Filter α} {s} + (hs : s ∈ ⨅ i, f i) : ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := + let ⟨_, _, V, hVs, _, _, hVU'⟩ := mem_iInf'.1 hs; ⟨V, hVs, hVU'⟩ + +theorem mem_iInf_of_finite {ι : Type*} [Finite ι] {α : Type*} {f : ι → Filter α} (s) : + (s ∈ ⨅ i, f i) ↔ ∃ t : ι → Set α, (∀ i, t i ∈ f i) ∧ s = ⋂ i, t i := by + refine ⟨exists_iInter_of_mem_iInf, ?_⟩ + rintro ⟨t, ht, rfl⟩ + exact iInter_mem.2 fun i => mem_iInf_of_mem i (ht i) + +/-! ### Lattice equations -/ + +theorem _root_.Pairwise.exists_mem_filter_of_disjoint {ι : Type*} [Finite ι] {l : ι → Filter α} + (hd : Pairwise (Disjoint on l)) : + ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ Pairwise (Disjoint on s) := by + have : Pairwise fun i j => ∃ (s : {s // s ∈ l i}) (t : {t // t ∈ l j}), Disjoint s.1 t.1 := by + simpa only [Pairwise, Function.onFun, Filter.disjoint_iff, exists_prop, Subtype.exists] using hd + choose! s t hst using this + refine ⟨fun i => ⋂ j, @s i j ∩ @t j i, fun i => ?_, fun i j hij => ?_⟩ + exacts [iInter_mem.2 fun j => inter_mem (@s i j).2 (@t j i).2, + (hst hij).mono ((iInter_subset _ j).trans inter_subset_left) + ((iInter_subset _ i).trans inter_subset_right)] + +theorem _root_.Set.PairwiseDisjoint.exists_mem_filter {ι : Type*} {l : ι → Filter α} {t : Set ι} + (hd : t.PairwiseDisjoint l) (ht : t.Finite) : + ∃ s : ι → Set α, (∀ i, s i ∈ l i) ∧ t.PairwiseDisjoint s := by + haveI := ht.to_subtype + rcases (hd.subtype _ _).exists_mem_filter_of_disjoint with ⟨s, hsl, hsd⟩ + lift s to (i : t) → {s // s ∈ l i} using hsl + rcases @Subtype.exists_pi_extension ι (fun i => { s // s ∈ l i }) _ _ s with ⟨s, rfl⟩ + exact ⟨fun i => s i, fun i => (s i).2, hsd.set_of_subtype _ _⟩ + + +theorem iInf_sets_eq_finite {ι : Type*} (f : ι → Filter α) : + (⨅ i, f i).sets = ⋃ t : Finset ι, (⨅ i ∈ t, f i).sets := by + rw [iInf_eq_iInf_finset, iInf_sets_eq] + exact directed_of_isDirected_le fun _ _ => biInf_mono + +theorem iInf_sets_eq_finite' (f : ι → Filter α) : + (⨅ i, f i).sets = ⋃ t : Finset (PLift ι), (⨅ i ∈ t, f (PLift.down i)).sets := by + rw [← iInf_sets_eq_finite, ← Equiv.plift.surjective.iInf_comp, Equiv.plift_apply] + +theorem mem_iInf_finite {ι : Type*} {f : ι → Filter α} (s) : + s ∈ iInf f ↔ ∃ t : Finset ι, s ∈ ⨅ i ∈ t, f i := + (Set.ext_iff.1 (iInf_sets_eq_finite f) s).trans mem_iUnion + +theorem mem_iInf_finite' {f : ι → Filter α} (s) : + s ∈ iInf f ↔ ∃ t : Finset (PLift ι), s ∈ ⨅ i ∈ t, f (PLift.down i) := + (Set.ext_iff.1 (iInf_sets_eq_finite' f) s).trans mem_iUnion + +/-- The dual version does not hold! `Filter α` is not a `CompleteDistribLattice`. -/ +-- See note [reducible non-instances] +abbrev coframeMinimalAxioms : Coframe.MinimalAxioms (Filter α) := + { Filter.instCompleteLatticeFilter with + iInf_sup_le_sup_sInf := fun f s t ⟨h₁, h₂⟩ => by + classical + rw [iInf_subtype'] + rw [sInf_eq_iInf', ← Filter.mem_sets, iInf_sets_eq_finite, mem_iUnion] at h₂ + obtain ⟨u, hu⟩ := h₂ + rw [← Finset.inf_eq_iInf] at hu + suffices ⨅ i : s, f ⊔ ↑i ≤ f ⊔ u.inf fun i => ↑i from this ⟨h₁, hu⟩ + refine Finset.induction_on u (le_sup_of_le_right le_top) ?_ + rintro ⟨i⟩ u _ ih + rw [Finset.inf_insert, sup_inf_left] + exact le_inf (iInf_le _ _) ih } + +instance instCoframe : Coframe (Filter α) := .ofMinimalAxioms coframeMinimalAxioms + +theorem mem_iInf_finset {s : Finset α} {f : α → Filter β} {t : Set β} : + (t ∈ ⨅ a ∈ s, f a) ↔ ∃ p : α → Set β, (∀ a ∈ s, p a ∈ f a) ∧ t = ⋂ a ∈ s, p a := by + classical + simp only [← Finset.set_biInter_coe, biInter_eq_iInter, iInf_subtype'] + refine ⟨fun h => ?_, ?_⟩ + · rcases (mem_iInf_of_finite _).1 h with ⟨p, hp, rfl⟩ + refine ⟨fun a => if h : a ∈ s then p ⟨a, h⟩ else univ, + fun a ha => by simpa [ha] using hp ⟨a, ha⟩, ?_⟩ + refine iInter_congr_of_surjective id surjective_id ?_ + rintro ⟨a, ha⟩ + simp [ha] + · rintro ⟨p, hpf, rfl⟩ + exact iInter_mem.2 fun a => mem_iInf_of_mem a (hpf a a.2) + + +@[elab_as_elim] +theorem iInf_sets_induct {f : ι → Filter α} {s : Set α} (hs : s ∈ iInf f) {p : Set α → Prop} + (uni : p univ) (ins : ∀ {i s₁ s₂}, s₁ ∈ f i → p s₂ → p (s₁ ∩ s₂)) : p s := by + classical + rw [mem_iInf_finite'] at hs + simp only [← Finset.inf_eq_iInf] at hs + rcases hs with ⟨is, his⟩ + induction is using Finset.induction_on generalizing s with + | empty => rwa [mem_top.1 his] + | insert _ ih => + rw [Finset.inf_insert, mem_inf_iff] at his + rcases his with ⟨s₁, hs₁, s₂, hs₂, rfl⟩ + exact ins hs₁ (ih hs₂) + +/-! #### `principal` equations -/ + +@[simp] +theorem iInf_principal_finset {ι : Type w} (s : Finset ι) (f : ι → Set α) : + ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by + classical + induction' s using Finset.induction_on with i s _ hs + · simp + · rw [Finset.iInf_insert, Finset.set_biInter_insert, hs, inf_principal] + +theorem iInf_principal {ι : Sort w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := by + cases nonempty_fintype (PLift ι) + rw [← iInf_plift_down, ← iInter_plift_down] + simpa using iInf_principal_finset Finset.univ (f <| PLift.down ·) + +/-- A special case of `iInf_principal` that is safe to mark `simp`. -/ +@[simp] +theorem iInf_principal' {ι : Type w} [Finite ι] (f : ι → Set α) : ⨅ i, 𝓟 (f i) = 𝓟 (⋂ i, f i) := + iInf_principal _ + +theorem iInf_principal_finite {ι : Type w} {s : Set ι} (hs : s.Finite) (f : ι → Set α) : + ⨅ i ∈ s, 𝓟 (f i) = 𝓟 (⋂ i ∈ s, f i) := by + lift s to Finset ι using hs + exact mod_cast iInf_principal_finset s f + +end Lattice + +/-! ### Eventually -/ + +@[simp] +theorem eventually_all {ι : Sort*} [Finite ι] {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i, p i x) ↔ ∀ i, ∀ᶠ x in l, p i x := by + simpa only [Filter.Eventually, setOf_forall] using iInter_mem + +@[simp] +theorem eventually_all_finite {ι} {I : Set ι} (hI : I.Finite) {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := by + simpa only [Filter.Eventually, setOf_forall] using biInter_mem hI + +alias _root_.Set.Finite.eventually_all := eventually_all_finite + +-- attribute [protected] Set.Finite.eventually_all + +@[simp] theorem eventually_all_finset {ι} (I : Finset ι) {l} {p : ι → α → Prop} : + (∀ᶠ x in l, ∀ i ∈ I, p i x) ↔ ∀ i ∈ I, ∀ᶠ x in l, p i x := + I.finite_toSet.eventually_all + +alias _root_.Finset.eventually_all := eventually_all_finset + +-- attribute [protected] Finset.eventually_all + +theorem eventually_imp_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∀ᶠ x in f, p → q x) ↔ p → ∀ᶠ x in f, q x := + eventually_all + + +/-! ### Frequently -/ + +@[simp] +theorem frequently_and_distrib_left {f : Filter α} {p : Prop} {q : α → Prop} : + (∃ᶠ x in f, p ∧ q x) ↔ p ∧ ∃ᶠ x in f, q x := by + simp only [Filter.Frequently, not_and, eventually_imp_distrib_left, Classical.not_imp] + +@[simp] +theorem frequently_and_distrib_right {f : Filter α} {p : α → Prop} {q : Prop} : + (∃ᶠ x in f, p x ∧ q) ↔ (∃ᶠ x in f, p x) ∧ q := by + simp only [@and_comm _ q, frequently_and_distrib_left] + +/-! +### Relation “eventually equal” +-/ + +section EventuallyEq +variable {l : Filter α} {f g : α → β} + +variable {l : Filter α} + +protected lemma EventuallyLE.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋃ i, s i) ≤ᶠ[l] ⋃ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ + let ⟨i, hi⟩ := mem_iUnion.1 hx'; mem_iUnion.2 ⟨i, hx i hi⟩ + +protected lemma EventuallyEq.iUnion [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋃ i, s i) =ᶠ[l] ⋃ i, t i := + (EventuallyLE.iUnion fun i ↦ (h i).le).antisymm <| .iUnion fun i ↦ (h i).symm.le + +protected lemma EventuallyLE.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i ≤ᶠ[l] t i) : (⋂ i, s i) ≤ᶠ[l] ⋂ i, t i := + (eventually_all.2 h).mono fun _x hx hx' ↦ mem_iInter.2 fun i ↦ hx i (mem_iInter.1 hx' i) + +protected lemma EventuallyEq.iInter [Finite ι] {s t : ι → Set α} + (h : ∀ i, s i =ᶠ[l] t i) : (⋂ i, s i) =ᶠ[l] ⋂ i, t i := + (EventuallyLE.iInter fun i ↦ (h i).le).antisymm <| .iInter fun i ↦ (h i).symm.le + +lemma _root_.Set.Finite.eventuallyLE_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := by + have := hs.to_subtype + rw [biUnion_eq_iUnion, biUnion_eq_iUnion] + exact .iUnion fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biUnion := Set.Finite.eventuallyLE_iUnion + +lemma _root_.Set.Finite.eventuallyEq_iUnion {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + (EventuallyLE.biUnion hs fun i hi ↦ (heq i hi).le).antisymm <| + .biUnion hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biUnion := Set.Finite.eventuallyEq_iUnion + +lemma _root_.Set.Finite.eventuallyLE_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := by + have := hs.to_subtype + rw [biInter_eq_iInter, biInter_eq_iInter] + exact .iInter fun i ↦ hle i.1 i.2 + +alias EventuallyLE.biInter := Set.Finite.eventuallyLE_iInter + +lemma _root_.Set.Finite.eventuallyEq_iInter {ι : Type*} {s : Set ι} (hs : s.Finite) + {f g : ι → Set α} (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + (EventuallyLE.biInter hs fun i hi ↦ (heq i hi).le).antisymm <| + .biInter hs fun i hi ↦ (heq i hi).symm.le + +alias EventuallyEq.biInter := Set.Finite.eventuallyEq_iInter + +lemma _root_.Finset.eventuallyLE_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋃ i ∈ s, f i) ≤ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iUnion {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋃ i ∈ s, f i) =ᶠ[l] (⋃ i ∈ s, g i) := + .biUnion s.finite_toSet heq + +lemma _root_.Finset.eventuallyLE_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (hle : ∀ i ∈ s, f i ≤ᶠ[l] g i) : (⋂ i ∈ s, f i) ≤ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet hle + +lemma _root_.Finset.eventuallyEq_iInter {ι : Type*} (s : Finset ι) {f g : ι → Set α} + (heq : ∀ i ∈ s, f i =ᶠ[l] g i) : (⋂ i ∈ s, f i) =ᶠ[l] (⋂ i ∈ s, g i) := + .biInter s.finite_toSet heq + +end EventuallyEq + +end Filter + +open Filter diff --git a/Mathlib/Order/Filter/Germ/Basic.lean b/Mathlib/Order/Filter/Germ/Basic.lean index 00156fd9037ce..c9ad7754c4946 100644 --- a/Mathlib/Order/Filter/Germ/Basic.lean +++ b/Mathlib/Order/Filter/Germ/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov, Abhimanyu Pallavi Sudhir -/ import Mathlib.Algebra.Module.Pi +import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Data.Int.Cast.Lemmas import Mathlib.Order.Filter.Tendsto diff --git a/Mathlib/Order/Filter/Subsingleton.lean b/Mathlib/Order/Filter/Subsingleton.lean index e079522700c04..c8d0816c3f17b 100644 --- a/Mathlib/Order/Filter/Subsingleton.lean +++ b/Mathlib/Order/Filter/Subsingleton.lean @@ -3,7 +3,6 @@ Copyright (c) 2023 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ -import Mathlib.Order.Filter.Bases import Mathlib.Order.Filter.Ultrafilter /-! # Subsingleton filters diff --git a/Mathlib/SetTheory/Cardinal/CountableCover.lean b/Mathlib/SetTheory/Cardinal/CountableCover.lean index 16eb4b46b329c..b3a0fe7e59ed2 100644 --- a/Mathlib/SetTheory/Cardinal/CountableCover.lean +++ b/Mathlib/SetTheory/Cardinal/CountableCover.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ import Mathlib.SetTheory.Cardinal.Arithmetic -import Mathlib.Order.Filter.Basic +import Mathlib.Order.Filter.Finite /-! # Cardinality of a set with a countable cover From 56428e8ac3a752358608d5efe49acdf7da3bf885 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 09:04:46 +0000 Subject: [PATCH 122/250] fix(CI): checkout the script in the action (#19449) Previously, the script was failing, since the script was unavailable. --- .github/workflows/maintainer_bors.yml | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 290ea7c66fa1b..67e3596b6a67e 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -101,6 +101,12 @@ jobs: python -m pip install --upgrade pip pip install zulip + - uses: actions/checkout@v4 + with: + ref: master + sparse-checkout: | + scripts/zulip_emoji_merge_delegate.py + - name: Run Zulip Emoji Merge Delegate Script if: ${{ ! steps.merge_or_delegate.outputs.mOrD == '' && ( steps.user_permission.outputs.require-result == 'true' || From 31eae5146f7f5b2884a2a96e50b93caa2932cd61 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 09:04:47 +0000 Subject: [PATCH 123/250] chore(discover-lean-pr-testing): only run on lean-toolchain changes (#19450) --- .github/workflows/discover-lean-pr-testing.yml | 2 ++ 1 file changed, 2 insertions(+) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 02e087dfb6085..349945a0ee8fe 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -4,6 +4,8 @@ on: push: branches: - nightly-testing + paths: + - lean-toolchain jobs: discover-lean-pr-testing: From 5fde9dcf705fe1d732a15da3dda14a624a95237f Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 09:04:48 +0000 Subject: [PATCH 124/250] fix: escape line breaks in create adaptation PR (#19451) --- scripts/create-adaptation-pr.sh | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/create-adaptation-pr.sh b/scripts/create-adaptation-pr.sh index b8ae7cf37b28f..f8b5deb203b7e 100755 --- a/scripts/create-adaptation-pr.sh +++ b/scripts/create-adaptation-pr.sh @@ -198,7 +198,7 @@ if git diff --name-only bump/$BUMPVERSION bump/nightly-$NIGHTLYDATE | grep -q .; echo "### [auto] post a link to the PR on Zulip" zulip_title="#$pr_number adaptations for nightly-$NIGHTLYDATE" - zulip_body="> $pr_title #$pr_number\\\n\\\nPlease review this PR. At the end of the month this diff will land in \\\`master\\\`." + zulip_body="> $pr_title #$pr_number"$'\n\nPlease review this PR. At the end of the month this diff will land in `master`.' echo "Posting the link to the PR in a new thread on the #nightly-testing channel on Zulip" echo "Here is the message:" From 5f4575362b440d45be2cc96e0654bf03a6a0e479 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 25 Nov 2024 09:33:01 +0000 Subject: [PATCH 125/250] refactor(Algebra/Category): turn homs in `AlgebraCat` into a structure (#19065) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit In the current design of concrete categories in mathlib, there are two sources of `erw`s: - Def-Eq abuse of `A ⟶ B` and `A →ₐ [R] B`: The type of `A ⟶ B` is by definition `A →ₐ[R] B`, but not at reducible transparency. So it happens that one `rw` lemma transforms a term in a form, where the next `rw` expects `A →ₐ B`. By supplying explicit `show` lines (copying the output of the pretty-printer), Lean correctly re-synthesizes the type to be `A →ₐ B` and the `rw` succeeds. This def-eq abuse hence always causes issues in the case where a lemma from the non-category theoretic part of the library should be applied. An insightful example is the following proof excerpt from current master (proof of `AlgebraCat.adj`): ```lean import Mathlib.Algebra.Category.AlgebraCat.Basic open CategoryTheory AlgebraCat variable {R : Type u} [CommRing R] set_option pp.analyze true example : free.{u} R ⊣ forget (AlgebraCat.{u} R) := Adjunction.mkOfHomEquiv { homEquiv := fun _ _ => (FreeAlgebra.lift _).symm homEquiv_naturality_left_symm := by intro X Y Z f g apply FreeAlgebra.hom_ext ext x simp only [Equiv.symm_symm, Function.comp_apply, free_map] /- Eq (α := Z) (((FreeAlgebra.lift R) (f ≫ g) : Prefunctor.obj (Functor.toPrefunctor (free R)) X ⟶ Z) (FreeAlgebra.ι R x) : ↑Z) _ -/ erw [FreeAlgebra.lift_ι_apply] sorry homEquiv_naturality_right := sorry } ``` The `simp` lemma `FreeAlgebra.lift_ι_apply` expects an `AlgHom`, but as the `pp.analyze` shows, the type is synthesized as a `⟶`. With a show line before the `erw`, Lean correctly synthesizes the type as an `AlgHom` again and the `rw` succeeds. The solution is to strictly distinguish between `A ⟶ B` and `A →ₐ[R] B`: We define a new `AlgebraCat.Hom` structure with a single field `hom : A →ₐ[R] B`. The above proof is mostly solved by `ext; simp` then, except for the `ext` lemma `FreeAlgebra.hom_ext` not applying. This is because now the type is `AlgebraCat.of R (FreeAlgebra R X) →ₐ[R] _` and Lean can't see through the `AlgebraCat.of` at reducible transparency. Hence the solution is to make `AlgebraCat.of` an `abbrev`. Finally, the proof is `by ext; simp` or even `by aesop`. - `FunLike`: An intermediate design adapted the changes from the first point, but with keeping a `FunLike` and an `AlgHomClass` instance on `A ⟶ B`. This lead to many additional `coe` lemmas for composition of morphisms, where it mattered which of the three appearing terms of `AlgebraCat` where of the form `AlgebraCat.of R A`. Eric Wieser suggested to replace the `FunLike` by a `CoeFun` which is inlined, so an invocation of `f x` turns into `f.hom x`. This has then worked very smoothly. So the key points are: - Homs in concrete categories should be one-field structures to maintain a strict type distinction. - Use `CoeFun` instead of `FunLike`, since the latter is automatically inlined. - Make `.of` constructors an `abbrev` to obtain the def-eq `↑(AlgebraCat.of R A) = A` at reducible transparency. For more details, see the [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Fixing.20the.20.60erw.60.20hell.20around.20concrete.20categories). --- .../Algebra/Category/AlgebraCat/Basic.lean | 218 +++++++++--------- .../Algebra/Category/AlgebraCat/Limits.lean | 52 ++--- .../Algebra/Category/AlgebraCat/Monoidal.lean | 2 +- .../Algebra/Category/BialgebraCat/Basic.lean | 4 +- .../Monoidal/Internal/Module.lean | 19 +- .../CliffordAlgebra/CategoryTheory.lean | 8 +- .../ConcreteCategory/AlgebraCat.lean | 53 +++++ 7 files changed, 195 insertions(+), 161 deletions(-) create mode 100644 MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean diff --git a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean index 9c06d5e619d85..1053dac09d757 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Basic.lean @@ -16,7 +16,6 @@ with the forgetful functors to `RingCat` and `ModuleCat`. We furthermore show th associating to a type the free `R`-algebra on that type is left adjoint to the forgetful functor. -/ - open CategoryTheory Limits universe v u @@ -25,6 +24,8 @@ variable (R : Type u) [CommRing R] /-- The category of R-algebras and their morphisms. -/ structure AlgebraCat where + private mk :: + /-- The underlying type. -/ carrier : Type v [isRing : Ring carrier] [isAlgebra : Algebra R carrier] @@ -46,22 +47,97 @@ instance : CoeSort (AlgebraCat R) (Type v) := attribute [coe] AlgebraCat.carrier +/-- The object in the category of R-algebras associated to a type equipped with the appropriate +typeclasses. This is the preferred way to construct a term of `AlgebraCat R`. -/ +abbrev of (X : Type v) [Ring X] [Algebra R X] : AlgebraCat.{v} R := + ⟨X⟩ + +lemma coe_of (X : Type v) [Ring X] [Algebra R X] : (of R X : Type v) = X := + rfl + +variable {R} in +/-- The type of morphisms in `AlgebraCat R`. -/ +@[ext] +structure Hom (A B : AlgebraCat.{v} R) where + private mk :: + /-- The underlying algebra map. -/ + hom : A →ₐ[R] B + instance : Category (AlgebraCat.{v} R) where - Hom A B := A →ₐ[R] B - id A := AlgHom.id R A - comp f g := g.comp f + Hom A B := Hom A B + id A := ⟨AlgHom.id R A⟩ + comp f g := ⟨g.hom.comp f.hom⟩ + +instance {M N : AlgebraCat.{v} R} : CoeFun (M ⟶ N) (fun _ ↦ M → N) where + coe f := f.hom -instance {M N : AlgebraCat.{v} R} : FunLike (M ⟶ N) M N := - AlgHom.funLike +@[simp] +lemma hom_id {A : AlgebraCat.{v} R} : (𝟙 A : A ⟶ A).hom = AlgHom.id R A := rfl + +/- Provided for rewriting. -/ +lemma id_apply (A : AlgebraCat.{v} R) (a : A) : + (𝟙 A : A ⟶ A) a = a := by simp + +@[simp] +lemma hom_comp {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) : + (f ≫ g).hom = g.hom.comp f.hom := rfl + +/- Provided for rewriting. -/ +lemma comp_apply {A B C : AlgebraCat.{v} R} (f : A ⟶ B) (g : B ⟶ C) (a : A) : + (f ≫ g) a = g (f a) := by simp + +@[ext] +lemma hom_ext {A B : AlgebraCat.{v} R} {f g : A ⟶ B} (hf : f.hom = g.hom) : f = g := + Hom.ext hf + +/-- Typecheck an `AlgHom` as a morphism in `AlgebraCat R`. -/ +abbrev ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] + (f : X →ₐ[R] Y) : of R X ⟶ of R Y := + ⟨f⟩ -instance {M N : AlgebraCat.{v} R} : AlgHomClass (M ⟶ N) R M N := - AlgHom.algHomClass +lemma hom_ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] + [Algebra R Y] (f : X →ₐ[R] Y) : (ofHom f).hom = f := rfl + +@[simp] +lemma ofHom_hom {A B : AlgebraCat.{v} R} (f : A ⟶ B) : + ofHom (Hom.hom f) = f := rfl + +@[simp] +lemma ofHom_id {X : Type v} [Ring X] [Algebra R X] : ofHom (AlgHom.id R X) = 𝟙 (of R X) := rfl + +@[simp] +lemma ofHom_comp {X Y Z : Type v} [Ring X] [Ring Y] [Ring Z] [Algebra R X] [Algebra R Y] + [Algebra R Z] (f : X →ₐ[R] Y) (g : Y →ₐ[R] Z) : + ofHom (g.comp f) = ofHom f ≫ ofHom g := + rfl + +lemma ofHom_apply {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] + [Algebra R Y] (f : X →ₐ[R] Y) (x : X) : ofHom f x = f x := rfl + +@[simp] +lemma inv_hom_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : A) : e.inv (e.hom x) = x := by + rw [← comp_apply] + simp + +@[simp] +lemma hom_inv_apply {A B : AlgebraCat.{v} R} (e : A ≅ B) (x : B) : e.hom (e.inv x) = x := by + rw [← comp_apply] + simp + +instance : Inhabited (AlgebraCat R) := + ⟨of R R⟩ instance : ConcreteCategory.{v} (AlgebraCat.{v} R) where forget := { obj := fun R => R - map := fun f => f.toFun } - forget_faithful := ⟨fun h => AlgHom.ext (by intros x; dsimp at h; rw [h])⟩ + map := fun f => f.hom } + forget_faithful := ⟨fun h => by ext x; simpa using congrFun h x⟩ + +lemma forget_obj {A : AlgebraCat.{v} R} : (forget (AlgebraCat.{v} R)).obj A = A := rfl + +lemma forget_map {A B : AlgebraCat.{v} R} (f : A ⟶ B) : + (forget (AlgebraCat.{v} R)).map f = f := + rfl instance {S : AlgebraCat.{v} R} : Ring ((forget (AlgebraCat R)).obj S) := (inferInstance : Ring S.carrier) @@ -72,12 +148,12 @@ instance {S : AlgebraCat.{v} R} : Algebra R ((forget (AlgebraCat R)).obj S) := instance hasForgetToRing : HasForget₂ (AlgebraCat.{v} R) RingCat.{v} where forget₂ := { obj := fun A => RingCat.of A - map := fun f => RingCat.ofHom f.toRingHom } + map := fun f => RingCat.ofHom f.hom.toRingHom } instance hasForgetToModule : HasForget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R) where forget₂ := { obj := fun M => ModuleCat.of R M - map := fun f => ModuleCat.asHom f.toLinearMap } + map := fun f => ModuleCat.asHom f.hom.toLinearMap } @[simp] lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @@ -86,33 +162,10 @@ lemma forget₂_module_obj (X : AlgebraCat.{v} R) : @[simp] lemma forget₂_module_map {X Y : AlgebraCat.{v} R} (f : X ⟶ Y) : - (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.toLinearMap := - rfl - -/-- The object in the category of R-algebras associated to a type equipped with the appropriate -typeclasses. -/ -def of (X : Type v) [Ring X] [Algebra R X] : AlgebraCat.{v} R := - ⟨X⟩ - -/-- Typecheck a `AlgHom` as a morphism in `AlgebraCat R`. -/ -def ofHom {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] - (f : X →ₐ[R] Y) : of R X ⟶ of R Y := - f - -@[simp] -theorem ofHom_apply {R : Type u} [CommRing R] {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] - [Algebra R Y] (f : X →ₐ[R] Y) (x : X) : ofHom f x = f x := + (forget₂ (AlgebraCat.{v} R) (ModuleCat.{v} R)).map f = ModuleCat.asHom f.hom.toLinearMap := rfl -instance : Inhabited (AlgebraCat R) := - ⟨of R R⟩ - -@[simp] -theorem coe_of (X : Type u) [Ring X] [Algebra R X] : (of R X : Type u) = X := - rfl - -variable {R} - +variable {R} in /-- Forgetting to the underlying type and then building the bundled object returns the original algebra. -/ @[simps] @@ -120,58 +173,20 @@ def ofSelfIso (M : AlgebraCat.{v} R) : AlgebraCat.of R M ≅ M where hom := 𝟙 M inv := 𝟙 M -variable {M N U : ModuleCat.{v} R} - -@[simp] -theorem id_apply (m : M) : (𝟙 M : M → M) m = m := - rfl - -@[simp] -theorem coe_comp (f : M ⟶ N) (g : N ⟶ U) : (f ≫ g : M → U) = g ∘ f := - rfl - -variable (R) - /-- The "free algebra" functor, sending a type `S` to the free algebra on `S`. -/ -@[simps!] +@[simps! obj map] def free : Type u ⥤ AlgebraCat.{u} R where - obj S := - { carrier := FreeAlgebra R S - isRing := Algebra.semiringToRing R } - map f := FreeAlgebra.lift _ <| FreeAlgebra.ι _ ∘ f - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - map_id := by intro X; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; rfl - map_comp := by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - intros; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; ext1 - -- Porting node: this ↓ `erw` used to be handled by the `simp` below it - erw [CategoryTheory.coe_comp] - simp only [CategoryTheory.coe_comp, Function.comp_apply, types_comp_apply] - -- Porting node: this ↓ `erw` and `rfl` used to be handled by the `simp` above - erw [FreeAlgebra.lift_ι_apply, FreeAlgebra.lift_ι_apply] - rfl + obj S := of R (FreeAlgebra R S) + map f := ofHom <| FreeAlgebra.lift _ <| FreeAlgebra.ι _ ∘ f /-- The free/forget adjunction for `R`-algebras. -/ def adj : free.{u} R ⊣ forget (AlgebraCat.{u} R) := Adjunction.mkOfHomEquiv - { homEquiv := fun _ _ => (FreeAlgebra.lift _).symm - -- Relying on `obviously` to fill out these proofs is very slow :( - homEquiv_naturality_left_symm := by - -- Porting note (https://github.com/leanprover-community/mathlib4/pull/11041): `apply FreeAlgebra.hom_ext` was `ext1`. - intros; apply FreeAlgebra.hom_ext; simp only [FreeAlgebra.ι_comp_lift]; ext1 - simp only [free_map, Equiv.symm_symm, FreeAlgebra.lift_ι_apply, CategoryTheory.coe_comp, - Function.comp_apply, types_comp_apply] - -- Porting node: this ↓ `erw` and `rfl` used to be handled by the `simp` above - erw [FreeAlgebra.lift_ι_apply, CategoryTheory.comp_apply, FreeAlgebra.lift_ι_apply, - Function.comp_apply, FreeAlgebra.lift_ι_apply] - rfl - homEquiv_naturality_right := by - intros; ext - simp only [CategoryTheory.coe_comp, Function.comp_apply, - FreeAlgebra.lift_symm_apply, types_comp_apply] - -- Porting note: proof used to be done after this ↑ `simp`; added ↓ two lines - erw [FreeAlgebra.lift_symm_apply, FreeAlgebra.lift_symm_apply] - rfl } + { homEquiv := fun _ _ => + { toFun := fun f ↦ (FreeAlgebra.lift _).symm f.hom + invFun := fun f ↦ ofHom <| (FreeAlgebra.lift _) f + left_inv := fun f ↦ by aesop + right_inv := fun f ↦ by simp [forget_obj, forget_map] } } instance : (forget (AlgebraCat.{u} R)).IsRightAdjoint := (adj R).isRightAdjoint @@ -184,31 +199,19 @@ variable {X₁ X₂ : Type u} @[simps] def AlgEquiv.toAlgebraIso {g₁ : Ring X₁} {g₂ : Ring X₂} {m₁ : Algebra R X₁} {m₂ : Algebra R X₂} (e : X₁ ≃ₐ[R] X₂) : AlgebraCat.of R X₁ ≅ AlgebraCat.of R X₂ where - hom := (e : X₁ →ₐ[R] X₂) - inv := (e.symm : X₂ →ₐ[R] X₁) - hom_inv_id := by ext x; exact e.left_inv x - inv_hom_id := by ext x; exact e.right_inv x + hom := AlgebraCat.ofHom (e : X₁ →ₐ[R] X₂) + inv := AlgebraCat.ofHom (e.symm : X₂ →ₐ[R] X₁) namespace CategoryTheory.Iso /-- Build a `AlgEquiv` from an isomorphism in the category `AlgebraCat R`. -/ @[simps] def toAlgEquiv {X Y : AlgebraCat R} (i : X ≅ Y) : X ≃ₐ[R] Y := - { i.hom with + { i.hom.hom with toFun := i.hom invFun := i.inv - left_inv := fun x => by - -- Porting note: was `by tidy` - change (i.hom ≫ i.inv) x = x - simp only [hom_inv_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [id_apply] - right_inv := fun x => by - -- Porting note: was `by tidy` - change (i.inv ≫ i.hom) x = x - simp only [inv_hom_id] - -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 - erw [id_apply] } + left_inv := fun x ↦ by simp + right_inv := fun x ↦ by simp } end CategoryTheory.Iso @@ -223,18 +226,5 @@ def algEquivIsoAlgebraIso {X Y : Type u} [Ring X] [Ring Y] [Algebra R X] [Algebr instance AlgebraCat.forget_reflects_isos : (forget (AlgebraCat.{u} R)).ReflectsIsomorphisms where reflects {X Y} f _ := by let i := asIso ((forget (AlgebraCat.{u} R)).map f) - let e : X ≃ₐ[R] Y := { f, i.toEquiv with } + let e : X ≃ₐ[R] Y := { f.hom, i.toEquiv with } exact e.toAlgebraIso.isIso_hom - -/-! -`@[simp]` lemmas for `AlgHom.comp` and categorical identities. --/ - -@[simp] theorem AlgHom.comp_id_algebraCat - {R} [CommRing R] {G : AlgebraCat.{u} R} {H : Type u} [Ring H] [Algebra R H] (f : G →ₐ[R] H) : - f.comp (𝟙 G) = f := - Category.id_comp (AlgebraCat.ofHom f) -@[simp] theorem AlgHom.id_algebraCat_comp - {R} [CommRing R] {G : Type u} [Ring G] [Algebra R G] {H : AlgebraCat.{u} R} (f : G →ₐ[R] H) : - AlgHom.comp (𝟙 H) f = f := - Category.comp_id (AlgebraCat.ofHom f) diff --git a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean index 94f312fee5673..ea39419a2442e 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Limits.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Limits.lean @@ -40,7 +40,7 @@ instance algebraObj (j) : def sectionsSubalgebra : Subalgebra R (∀ j, F.obj j) := { SemiRingCat.sectionsSubsemiring (F ⋙ forget₂ (AlgebraCat R) RingCat.{w} ⋙ forget₂ RingCat SemiRingCat.{w}) with - algebraMap_mem' := fun r _ _ f => (F.map f).commutes r } + algebraMap_mem' := fun r _ _ f => (F.map f).hom.commutes r } instance (F : J ⥤ AlgebraCat.{w} R) : Ring (F ⋙ forget _).sections := inferInstanceAs <| Ring (sectionsSubalgebra F) @@ -88,9 +88,10 @@ namespace HasLimits def limitCone : Cone F where pt := AlgebraCat.of R (Types.Small.limitCone (F ⋙ forget _)).pt π := - { app := limitπAlgHom F - naturality := fun _ _ f => - AlgHom.coe_fn_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } + { app := fun j ↦ ofHom <| limitπAlgHom F j + naturality := fun _ _ f => by + ext : 1 + exact AlgHom.coe_fn_injective ((Types.Small.limitCone (F ⋙ forget _)).π.naturality f) } /-- Witness that the limit cone in `AlgebraCat R` is a limit cone. (Internal use only; use the limits API.) @@ -101,41 +102,36 @@ def limitConeIsLimit : IsLimit (limitCone.{v, w} F) := by -- Porting note: in mathlib3 the function term -- `fun v => ⟨fun j => ((forget (AlgebraCat R)).mapCone s).π.app j v` -- was provided by unification, and the last argument `(fun s => _)` was `(fun s => rfl)`. - (fun s => { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_, - commutes' := ?_ }) + (fun s => ofHom + { toFun := _, map_one' := ?_, map_mul' := ?_, map_zero' := ?_, map_add' := ?_, + commutes' := ?_ }) (fun s => rfl) · congr ext j - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app, - forget_map_eq_coe] - rw [map_one] - rfl + simp only [Functor.mapCone_π_app, forget_map, map_one, Pi.one_apply] · intro x y - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app] - rw [← equivShrink_mul] - apply congrArg ext j - simp only [Functor.comp_obj, Functor.mapCone_pt, Functor.mapCone_π_app, - forget_map_eq_coe, map_mul] - rfl - · simp only [Functor.mapCone_π_app, forget_map_eq_coe] - congr - funext j - simp only [map_zero, Pi.zero_apply] + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + Types.Small.limitCone_pt, equivShrink_symm_mul] + apply map_mul + · ext j + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + equivShrink_symm_zero] + apply map_zero · intro x y - simp only [Functor.mapCone_π_app] - rw [← equivShrink_add] - apply congrArg ext j - simp only [forget_map_eq_coe, map_add] - rfl + simp only [Functor.comp_obj, forget_obj, Equiv.toFun_as_coe, Functor.mapCone_pt, + Functor.mapCone_π_app, forget_map, Equiv.symm_apply_apply, + Types.Small.limitCone_pt, equivShrink_symm_add] + apply map_add · intro r - simp only [← Shrink.algEquiv_symm_apply _ R, limitCone, Equiv.algebraMap_def, - Equiv.symm_symm] + simp only [← Shrink.algEquiv_symm_apply _ R, limitCone, Equiv.algebraMap_def, Equiv.symm_symm] apply congrArg apply Subtype.ext ext j - exact (s.π.app j).commutes r + exact (s.π.app j).hom.commutes r end HasLimits diff --git a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean index ac5df5ca0f2dd..fa16a8aac4bcd 100644 --- a/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean +++ b/Mathlib/Algebra/Category/AlgebraCat/Monoidal.lean @@ -37,7 +37,7 @@ noncomputable abbrev tensorObj (X Y : AlgebraCat.{u} R) : AlgebraCat.{u} R := `AlgebraCat.instMonoidalCategory`. -/ noncomputable abbrev tensorHom {W X Y Z : AlgebraCat.{u} R} (f : W ⟶ X) (g : Y ⟶ Z) : tensorObj W Y ⟶ tensorObj X Z := - Algebra.TensorProduct.map f g + ofHom <| Algebra.TensorProduct.map f.hom g.hom open MonoidalCategory diff --git a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean index fb1f40ffe1018..59fe5b8fefba2 100644 --- a/Mathlib/Algebra/Category/BialgebraCat/Basic.lean +++ b/Mathlib/Algebra/Category/BialgebraCat/Basic.lean @@ -103,7 +103,7 @@ instance concreteCategory : ConcreteCategory.{v} (BialgebraCat.{v} R) where instance hasForgetToAlgebra : HasForget₂ (BialgebraCat R) (AlgebraCat R) where forget₂ := { obj := fun X => AlgebraCat.of R X - map := fun {X Y} f => (f.toBialgHom : X →ₐ[R] Y) } + map := fun {X Y} f => AlgebraCat.ofHom f.toBialgHom } @[simp] theorem forget₂_algebra_obj (X : BialgebraCat R) : @@ -112,7 +112,7 @@ theorem forget₂_algebra_obj (X : BialgebraCat R) : @[simp] theorem forget₂_algebra_map (X Y : BialgebraCat R) (f : X ⟶ Y) : - (forget₂ (BialgebraCat R) (AlgebraCat R)).map f = (f.toBialgHom : X →ₐ[R] Y) := + (forget₂ (BialgebraCat R) (AlgebraCat R)).map f = AlgebraCat.ofHom f.toBialgHom := rfl instance hasForgetToCoalgebra : HasForget₂ (BialgebraCat R) (CoalgebraCat R) where diff --git a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean index 5974753934292..9341953c58518 100644 --- a/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean +++ b/Mathlib/CategoryTheory/Monoidal/Internal/Module.lean @@ -88,7 +88,7 @@ theorem algebraMap (A : Mon_ (ModuleCat.{u} R)) (r : R) : algebraMap R A.X r = A @[simps!] def functor : Mon_ (ModuleCat.{u} R) ⥤ AlgebraCat R where obj A := AlgebraCat.of R A.X - map {_ _} f := + map {_ _} f := AlgebraCat.ofHom { f.hom.toAddMonoidHom with toFun := f.hom map_one' := LinearMap.congr_fun f.one_hom (1 : R) @@ -134,9 +134,7 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): `ext` did not pick up `TensorProduct.ext` refine TensorProduct.ext <| TensorProduct.ext <| LinearMap.ext fun x => LinearMap.ext fun y => LinearMap.ext fun z => ?_ - dsimp only [AlgebraCat.id_apply, TensorProduct.mk_apply, LinearMap.compr₂_apply, - Function.comp_apply, ModuleCat.MonoidalCategory.tensorHom_tmul, AlgebraCat.coe_comp, - MonoidalCategory.associator_hom_apply] + dsimp only [compr₂_apply, TensorProduct.mk_apply] rw [compr₂_apply, compr₂_apply] -- This used to be `rw`, but we need `erw` after https://github.com/leanprover/lean4/pull/2644 erw [CategoryTheory.comp_apply, @@ -152,9 +150,9 @@ def inverseObj (A : AlgebraCat.{u} R) : Mon_ (ModuleCat.{u} R) where def inverse : AlgebraCat.{u} R ⥤ Mon_ (ModuleCat.{u} R) where obj := inverseObj map f := - { hom := f.toLinearMap - one_hom := LinearMap.ext f.commutes - mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f } + { hom := f.hom.toLinearMap + one_hom := LinearMap.ext f.hom.commutes + mul_hom := TensorProduct.ext <| LinearMap.ext₂ <| map_mul f.hom } end MonModuleEquivalenceAlgebra @@ -193,14 +191,14 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where counitIso := NatIso.ofComponents (fun A => - { hom := + { hom := AlgebraCat.ofHom { toFun := _root_.id map_zero' := rfl map_add' := fun _ _ => rfl map_one' := (algebraMap R A).map_one map_mul' := fun x y => @LinearMap.mul'_apply R _ _ _ _ _ _ x y commutes' := fun _ => rfl } - inv := + inv := AlgebraCat.ofHom { toFun := _root_.id map_zero' := rfl map_add' := fun _ _ => rfl @@ -208,9 +206,6 @@ def monModuleEquivalenceAlgebra : Mon_ (ModuleCat.{u} R) ≌ AlgebraCat R where map_mul' := fun x y => (@LinearMap.mul'_apply R _ _ _ _ _ _ x y).symm commutes' := fun _ => rfl } }) --- These lemmas have always been bad (https://github.com/leanprover-community/mathlib4/issues/7657), but https://github.com/leanprover/lean4/pull/2644 made `simp` start noticing -attribute [nolint simpNF] ModuleCat.MonModuleEquivalenceAlgebra.functor_map_apply - /-- The equivalence `Mon_ (ModuleCat R) ≌ AlgebraCat R` is naturally compatible with the forgetful functors to `ModuleCat R`. -/ diff --git a/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean b/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean index 2311b81349cc4..0b5239188f120 100644 --- a/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean +++ b/Mathlib/LinearAlgebra/CliffordAlgebra/CategoryTheory.lean @@ -25,7 +25,7 @@ variable {R : Type u} [CommRing R] This is `CliffordAlgebra.map` through the lens of category theory. -/ @[simps] def QuadraticModuleCat.cliffordAlgebra : QuadraticModuleCat.{u} R ⥤ AlgebraCat.{u} R where - obj M := { carrier := CliffordAlgebra M.form } - map {_M _N} f := CliffordAlgebra.map f.toIsometry - map_id _M := CliffordAlgebra.map_id _ - map_comp {_M _N _P} f g := (CliffordAlgebra.map_comp_map g.toIsometry f.toIsometry).symm + obj M := AlgebraCat.of R (CliffordAlgebra M.form) + map {_M _N} f := AlgebraCat.ofHom <| CliffordAlgebra.map f.toIsometry + map_id _M := by simp + map_comp {_M _N _P} f g := by ext; simp diff --git a/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean b/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean new file mode 100644 index 0000000000000..406d12cc89de4 --- /dev/null +++ b/MathlibTest/CategoryTheory/ConcreteCategory/AlgebraCat.lean @@ -0,0 +1,53 @@ +import Mathlib + +universe v u + +open CategoryTheory AlgebraCat + +set_option maxHeartbeats 10000 +set_option synthInstance.maxHeartbeats 2000 + +variable (R : Type u) [CommRing R] + +/- We test if all the coercions and `map_add` lemmas trigger correctly. -/ + +example (X : Type u) [Ring X] [Algebra R X] : ⇑(𝟙 (of R X)) = id := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) : + ⇑(ofHom f) = ⇑f := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] (f : X →ₐ[R] Y) + (x : X) : (ofHom f) x = f x := by simp + +example {X Y Z : AlgebraCat R} (f : X ⟶ Y) (g : Y ⟶ Z) : ⇑(f ≫ g) = ⇑g ∘ ⇑f := by simp + +example {X Y Z : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] [Ring Z] + [Algebra R Z] (f : X →ₐ[R] Y) (g : Y →ₐ[R] Z) : + ⇑(ofHom f ≫ ofHom g) = g ∘ f := by simp + +example {X Y : Type v} [Ring X] [Algebra R X] [Ring Y] [Algebra R Y] {Z : AlgebraCat R} + (f : X →ₐ[R] Y) (g : of R Y ⟶ Z) : + ⇑(ofHom f ≫ g) = g ∘ f := by simp + +example {X Y : AlgebraCat R} {Z : Type v} [Ring Z] [Algebra R Z] (f : X ⟶ Y) (g : Y ⟶ of R Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {Y Z : AlgebraCat R} {X : Type v} [Ring X] [Algebra R X] (f : of R X ⟶ Y) (g : Y ⟶ Z) : + ⇑(f ≫ g) = g ∘ f := by simp + +example {X Y Z : AlgebraCat R} (f : X ⟶ Y) (g : Y ⟶ Z) (x : X) : (f ≫ g) x = g (f x) := by simp + +example {X Y : AlgebraCat R} (e : X ≅ Y) (x : X) : e.inv (e.hom x) = x := by simp + +example {X Y : AlgebraCat R} (e : X ≅ Y) (y : Y) : e.hom (e.inv y) = y := by simp + +example (X : AlgebraCat R) : ⇑(𝟙 X) = id := by simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) (x y : M) : f (x + y) = f x + f y := by + simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) : f 0 = 0 := by + simp + +example {M N : AlgebraCat.{v} R} (f : M ⟶ N) (r : R) (m : M) : f (r • m) = r • f m := by + simp From a5ae0c1c33f439193d3e3e967e514035edb9fe55 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 25 Nov 2024 09:33:02 +0000 Subject: [PATCH 126/250] chore: remove the `Smooth` alias for infinitely differentiable functions between manifolds (#19296) The abbreviation `Smooth` for `ContMDiff` with infinite smoothness had been introduced in mathlib in ancient times. This has lead to a big API duplication, for no visible gain. We deprecate this abbreviation (with its friends `SmoothAt` and so on) and all the lemmas involving it. In the longer run, we might consider reintroducing it as a mere notation (without any specific API), but this is not part of this PR, devoted purely to the deprecation. No mathematical content added. --- .../Complex/UpperHalfPlane/Manifold.lean | 16 +- .../Distribution/AEEqOfIntegralContDiff.lean | 10 +- .../Geometry/Manifold/Algebra/LieGroup.lean | 125 ++++------- Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 177 +++++++--------- .../Manifold/Algebra/SmoothFunctions.lean | 30 +-- .../Geometry/Manifold/Algebra/Structures.lean | 8 +- Mathlib/Geometry/Manifold/BumpFunction.lean | 21 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 4 +- .../Geometry/Manifold/ContMDiff/Basic.lean | 114 ++++------ Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 135 ++++-------- .../Manifold/ContMDiff/NormedSpace.lean | 27 +-- .../Geometry/Manifold/ContMDiff/Product.lean | 124 ++++------- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 11 +- Mathlib/Geometry/Manifold/ContMDiffMap.lean | 8 +- Mathlib/Geometry/Manifold/Diffeomorph.lean | 12 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 29 +-- .../Manifold/MFDeriv/UniqueDifferential.lean | 2 +- .../Geometry/Manifold/PartitionOfUnity.lean | 77 ++++--- .../Manifold/Sheaf/LocallyRingedSpace.lean | 2 +- Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 11 +- .../Geometry/Manifold/VectorBundle/Basic.lean | 198 ++++++++---------- .../VectorBundle/FiberwiseLinear.lean | 44 ++-- .../Geometry/Manifold/VectorBundle/Hom.lean | 21 +- .../Manifold/VectorBundle/Pullback.lean | 6 +- .../Manifold/VectorBundle/SmoothSection.lean | 10 +- .../Manifold/VectorBundle/Tangent.lean | 2 +- .../Geometry/Manifold/WhitneyEmbedding.lean | 13 +- 27 files changed, 495 insertions(+), 742 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean index c821c8feafdcb..9529dbcb764f2 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Manifold.lean @@ -26,15 +26,17 @@ instance : SmoothManifoldWithCorners 𝓘(ℂ) ℍ := UpperHalfPlane.isOpenEmbedding_coe.singleton_smoothManifoldWithCorners /-- The inclusion map `ℍ → ℂ` is a smooth map of manifolds. -/ -theorem smooth_coe : Smooth 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt +theorem contMDiff_coe : ContMDiff 𝓘(ℂ) 𝓘(ℂ) ⊤ ((↑) : ℍ → ℂ) := fun _ => contMDiffAt_extChartAt + +@[deprecated (since := "2024-11-20")] alias smooth_coe := contMDiff_coe /-- The inclusion map `ℍ → ℂ` is a differentiable map of manifolds. -/ theorem mdifferentiable_coe : MDifferentiable 𝓘(ℂ) 𝓘(ℂ) ((↑) : ℍ → ℂ) := - smooth_coe.mdifferentiable + contMDiff_coe.mdifferentiable (by simp) -lemma smoothAt_ofComplex {z : ℂ} (hz : 0 < z.im) : - SmoothAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := by - rw [SmoothAt, contMDiffAt_iff] +lemma contMDiffAt_ofComplex {z : ℂ} (hz : 0 < z.im) : + ContMDiffAt 𝓘(ℂ) 𝓘(ℂ) ⊤ ofComplex z := by + rw [contMDiffAt_iff] constructor · -- continuity at z rw [ContinuousAt, nhds_induced, tendsto_comap_iff] @@ -48,9 +50,11 @@ lemma smoothAt_ofComplex {z : ℂ} (hz : 0 < z.im) : Set.range_id, id_eq, contDiffWithinAt_univ] exact contDiffAt_id.congr_of_eventuallyEq (eventuallyEq_coe_comp_ofComplex hz) +@[deprecated (since := "2024-11-20")] alias smoothAt_ofComplex := contMDiffAt_ofComplex + lemma mdifferentiableAt_ofComplex {z : ℂ} (hz : 0 < z.im) : MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) ofComplex z := - (smoothAt_ofComplex hz).mdifferentiableAt + (contMDiffAt_ofComplex hz).mdifferentiableAt (by simp) lemma mdifferentiableAt_iff {f : ℍ → ℂ} {τ : ℍ} : MDifferentiableAt 𝓘(ℂ) 𝓘(ℂ) f τ ↔ DifferentiableAt ℂ (f ∘ ofComplex) ↑τ := by diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index fb4a816b271bc..b6d2c67574d34 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -40,7 +40,7 @@ variable {H : Type*} [TopologicalSpace H] (I : ModelWithCorners ℝ E H) when multiplied by any smooth compactly supported function, then `f` vanishes almost everywhere. -/ theorem ae_eq_zero_of_integral_smooth_smul_eq_zero [SigmaCompactSpace M] (hf : LocallyIntegrable f μ) - (h : ∀ g : M → ℝ, Smooth I 𝓘(ℝ) g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : + (h : ∀ g : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, f x = 0 := by -- record topological properties of `M` have := I.locallyCompactSpace @@ -60,7 +60,7 @@ theorem ae_eq_zero_of_integral_smooth_smul_eq_zero [SigmaCompactSpace M] let v : ℕ → Set M := fun n ↦ thickening (u n) s obtain ⟨K, K_compact, vK⟩ : ∃ K, IsCompact K ∧ ∀ n, v n ⊆ K := ⟨_, hδ, fun n ↦ thickening_subset_cthickening_of_le (u_pos n).2.le _⟩ - have : ∀ n, ∃ (g : M → ℝ), support g = v n ∧ Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 + have : ∀ n, ∃ (g : M → ℝ), support g = v n ∧ ContMDiff I 𝓘(ℝ) ⊤ g ∧ Set.range g ⊆ Set.Icc 0 1 ∧ ∀ x ∈ s, g x = 1 := by intro n rcases exists_msmooth_support_eq_eq_one_iff I isOpen_thickening hs.isClosed @@ -121,7 +121,7 @@ instance (U : Opens M) : BorelSpace U := inferInstanceAs (BorelSpace (U : Set M) nonrec theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero' {U : Set M} (hU : IsOpen U) (hSig : IsSigmaCompact U) (hf : LocallyIntegrableOn f U μ) (h : ∀ g : M → ℝ, - Smooth I 𝓘(ℝ) g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := by have meas_U := hU.measurableSet rw [← ae_restrict_iff' meas_U, ae_restrict_iff_subtype meas_U] @@ -146,7 +146,7 @@ variable [SigmaCompactSpace M] theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero {U : Set M} (hU : IsOpen U) (hf : LocallyIntegrableOn f U μ) (h : ∀ g : M → ℝ, - Smooth I 𝓘(ℝ) g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := haveI := I.locallyCompactSpace haveI := ChartedSpace.locallyCompactSpace H M @@ -160,7 +160,7 @@ theorem IsOpen.ae_eq_zero_of_integral_smooth_smul_eq_zero {U : Set M} (hU : IsOp when multiplied by any smooth compactly supported function, then they coincide almost everywhere. -/ theorem ae_eq_of_integral_smooth_smul_eq (hf : LocallyIntegrable f μ) (hf' : LocallyIntegrable f' μ) (h : ∀ (g : M → ℝ), - Smooth I 𝓘(ℝ) g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : + ContMDiff I 𝓘(ℝ) ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : ∀ᵐ x ∂μ, f x = f' x := by have : ∀ᵐ x ∂μ, (f - f') x = 0 := by apply ae_eq_zero_of_integral_smooth_smul_eq_zero I (hf.sub hf') diff --git a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean index c054be8efa60b..c9dd661d44afd 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LieGroup.lean @@ -58,7 +58,7 @@ class LieAddGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Top {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [AddGroup G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothAdd I G : Prop where /-- Negation is smooth in an additive Lie group. -/ - smooth_neg : Smooth I I fun a : G => -a + smooth_neg : ContMDiff I I ⊤ fun a : G => -a -- See note [Design choices about smooth algebraic structures] /-- A (multiplicative) Lie group is a group and a smooth manifold at the same time in which @@ -68,7 +68,7 @@ class LieGroup {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topolo {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Group G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothMul I G : Prop where /-- Inversion is smooth in a Lie group. -/ - smooth_inv : Smooth I I fun a : G => a⁻¹ + smooth_inv : ContMDiff I I ⊤ fun a : G => a⁻¹ /-! ### Smoothness of inversion, negation, division and subtraction @@ -91,28 +91,31 @@ variable (I) /-- In a Lie group, inversion is a smooth map. -/ @[to_additive "In an additive Lie group, inversion is a smooth map."] -theorem smooth_inv : Smooth I I fun x : G => x⁻¹ := +theorem contMDiff_inv : ContMDiff I I ⊤ fun x : G => x⁻¹ := LieGroup.smooth_inv +@[deprecated (since := "2024-11-21")] alias smooth_inv := contMDiff_inv +@[deprecated (since := "2024-11-21")] alias smooth_neg := contMDiff_neg + include I in /-- A Lie group is a topological group. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ @[to_additive "An additive Lie group is an additive topological group. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]."] theorem topologicalGroup_of_lieGroup : TopologicalGroup G := - { continuousMul_of_smooth I with continuous_inv := (smooth_inv I).continuous } + { continuousMul_of_smooth I with continuous_inv := (contMDiff_inv I).continuous } end @[to_additive] theorem ContMDiffWithinAt.inv {f : M → G} {s : Set M} {x₀ : M} (hf : ContMDiffWithinAt I' I n f s x₀) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s x₀ := - ((smooth_inv I).of_le le_top).contMDiffAt.contMDiffWithinAt.comp x₀ hf <| Set.mapsTo_univ _ _ + ((contMDiff_inv I).of_le le_top).contMDiffAt.contMDiffWithinAt.comp x₀ hf <| Set.mapsTo_univ _ _ @[to_additive] theorem ContMDiffAt.inv {f : M → G} {x₀ : M} (hf : ContMDiffAt I' I n f x₀) : ContMDiffAt I' I n (fun x => (f x)⁻¹) x₀ := - ((smooth_inv I).of_le le_top).contMDiffAt.comp x₀ hf + ((contMDiff_inv I).of_le le_top).contMDiffAt.comp x₀ hf @[to_additive] theorem ContMDiffOn.inv {f : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s) : @@ -122,24 +125,15 @@ theorem ContMDiffOn.inv {f : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s) theorem ContMDiff.inv {f : M → G} (hf : ContMDiff I' I n f) : ContMDiff I' I n fun x => (f x)⁻¹ := fun x => (hf x).inv -@[to_additive] -nonrec theorem SmoothWithinAt.inv {f : M → G} {s : Set M} {x₀ : M} - (hf : SmoothWithinAt I' I f s x₀) : SmoothWithinAt I' I (fun x => (f x)⁻¹) s x₀ := - hf.inv - -@[to_additive] -nonrec theorem SmoothAt.inv {f : M → G} {x₀ : M} (hf : SmoothAt I' I f x₀) : - SmoothAt I' I (fun x => (f x)⁻¹) x₀ := - hf.inv +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.inv := ContMDiffWithinAt.inv +@[deprecated (since := "2024-11-21")] alias SmoothAt.inv := ContMDiffAt.inv +@[deprecated (since := "2024-11-21")] alias SmoothOn.inv := ContMDiffOn.inv +@[deprecated (since := "2024-11-21")] alias Smooth.inv := ContMDiff.inv -@[to_additive] -nonrec theorem SmoothOn.inv {f : M → G} {s : Set M} (hf : SmoothOn I' I f s) : - SmoothOn I' I (fun x => (f x)⁻¹) s := - hf.inv - -@[to_additive] -nonrec theorem Smooth.inv {f : M → G} (hf : Smooth I' I f) : Smooth I' I fun x => (f x)⁻¹ := - hf.inv +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.neg := ContMDiffWithinAt.neg +@[deprecated (since := "2024-11-21")] alias SmoothAt.neg := ContMDiffAt.neg +@[deprecated (since := "2024-11-21")] alias SmoothOn.neg := ContMDiffOn.neg +@[deprecated (since := "2024-11-21")] alias Smooth.neg := ContMDiff.neg @[to_additive] theorem ContMDiffWithinAt.div {f g : M → G} {s : Set M} {x₀ : M} @@ -161,26 +155,15 @@ theorem ContMDiffOn.div {f g : M → G} {s : Set M} (hf : ContMDiffOn I' I n f s theorem ContMDiff.div {f g : M → G} (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) : ContMDiff I' I n fun x => f x / g x := by simp_rw [div_eq_mul_inv]; exact hf.mul hg.inv -@[to_additive] -nonrec theorem SmoothWithinAt.div {f g : M → G} {s : Set M} {x₀ : M} - (hf : SmoothWithinAt I' I f s x₀) (hg : SmoothWithinAt I' I g s x₀) : - SmoothWithinAt I' I (fun x => f x / g x) s x₀ := - hf.div hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div := ContMDiffWithinAt.div +@[deprecated (since := "2024-11-21")] alias SmoothAt.div := ContMDiffAt.div +@[deprecated (since := "2024-11-21")] alias SmoothOn.div := ContMDiffOn.div +@[deprecated (since := "2024-11-21")] alias Smooth.div := ContMDiff.div -@[to_additive] -nonrec theorem SmoothAt.div {f g : M → G} {x₀ : M} (hf : SmoothAt I' I f x₀) - (hg : SmoothAt I' I g x₀) : SmoothAt I' I (fun x => f x / g x) x₀ := - hf.div hg - -@[to_additive] -nonrec theorem SmoothOn.div {f g : M → G} {s : Set M} (hf : SmoothOn I' I f s) - (hg : SmoothOn I' I g s) : SmoothOn I' I (f / g) s := - hf.div hg - -@[to_additive] -nonrec theorem Smooth.div {f g : M → G} (hf : Smooth I' I f) (hg : Smooth I' I g) : - Smooth I' I (f / g) := - hf.div hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.sub := ContMDiffWithinAt.sub +@[deprecated (since := "2024-11-21")] alias SmoothAt.sub := ContMDiffAt.sub +@[deprecated (since := "2024-11-21")] alias SmoothOn.sub := ContMDiffOn.sub +@[deprecated (since := "2024-11-21")] alias Smooth.sub := ContMDiff.sub end PointwiseDivision @@ -195,7 +178,7 @@ instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] {H' : Type*} [TopologicalSpace H'] {I' : ModelWithCorners 𝕜 E' H'} {G' : Type*} [TopologicalSpace G'] [ChartedSpace H' G'] [Group G'] [LieGroup I' G'] : LieGroup (I.prod I') (G × G') := - { SmoothMul.prod _ _ _ _ with smooth_inv := smooth_fst.inv.prod_mk smooth_snd.inv } + { SmoothMul.prod _ _ _ _ with smooth_inv := contMDiff_fst.inv.prod_mk contMDiff_snd.inv } end Product @@ -219,7 +202,7 @@ class SmoothInv₀ {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [To {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Inv G] [Zero G] [TopologicalSpace G] [ChartedSpace H G] : Prop where /-- Inversion is smooth away from `0`. -/ - smoothAt_inv₀ : ∀ ⦃x : G⦄, x ≠ 0 → SmoothAt I I (fun y ↦ y⁻¹) x + smoothAt_inv₀ : ∀ ⦃x : G⦄, x ≠ 0 → ContMDiffAt I I ⊤ (fun y ↦ y⁻¹) x instance {𝕜 : Type*} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] : SmoothInv₀ 𝓘(𝕜) 𝕜 := { smoothAt_inv₀ := by @@ -235,28 +218,33 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS {I' : ModelWithCorners 𝕜 E' H'} {M : Type*} [TopologicalSpace M] [ChartedSpace H' M] {n : ℕ∞} {f : M → G} -theorem smoothAt_inv₀ {x : G} (hx : x ≠ 0) : SmoothAt I I (fun y ↦ y⁻¹) x := +theorem contMDiffAt_inv₀ {x : G} (hx : x ≠ 0) : ContMDiffAt I I ⊤ (fun y ↦ y⁻¹) x := SmoothInv₀.smoothAt_inv₀ hx +@[deprecated (since := "2024-11-21")] alias smoothAt_inv₀ := contMDiffAt_inv₀ + include I in /-- In a manifold with smooth inverse away from `0`, the inverse is continuous away from `0`. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ theorem hasContinuousInv₀_of_hasSmoothInv₀ : HasContinuousInv₀ G := - { continuousAt_inv₀ := fun _ hx ↦ (smoothAt_inv₀ I hx).continuousAt } + { continuousAt_inv₀ := fun _ hx ↦ (contMDiffAt_inv₀ I hx).continuousAt } -theorem SmoothOn_inv₀ : SmoothOn I I (Inv.inv : G → G) {0}ᶜ := fun _x hx => - (smoothAt_inv₀ I hx).smoothWithinAt +theorem contMDiffOn_inv₀ : ContMDiffOn I I ⊤ (Inv.inv : G → G) {0}ᶜ := fun _x hx => + (contMDiffAt_inv₀ I hx).contMDiffWithinAt + +@[deprecated (since := "2024-11-21")] alias smoothOn_inv₀ := contMDiffOn_inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn_inv₀ := contMDiffOn_inv₀ variable {I} {s : Set M} {a : M} theorem ContMDiffWithinAt.inv₀ (hf : ContMDiffWithinAt I' I n f s a) (ha : f a ≠ 0) : ContMDiffWithinAt I' I n (fun x => (f x)⁻¹) s a := - (smoothAt_inv₀ I ha).contMDiffAt.comp_contMDiffWithinAt a hf + ((contMDiffAt_inv₀ I ha).of_le le_top).comp_contMDiffWithinAt a hf theorem ContMDiffAt.inv₀ (hf : ContMDiffAt I' I n f a) (ha : f a ≠ 0) : ContMDiffAt I' I n (fun x ↦ (f x)⁻¹) a := - (smoothAt_inv₀ I ha).contMDiffAt.comp a hf + ((contMDiffAt_inv₀ I ha).of_le le_top).comp a hf theorem ContMDiff.inv₀ (hf : ContMDiff I' I n f) (h0 : ∀ x, f x ≠ 0) : ContMDiff I' I n (fun x ↦ (f x)⁻¹) := @@ -266,20 +254,10 @@ theorem ContMDiffOn.inv₀ (hf : ContMDiffOn I' I n f s) (h0 : ∀ x ∈ s, f x ContMDiffOn I' I n (fun x => (f x)⁻¹) s := fun x hx ↦ ContMDiffWithinAt.inv₀ (hf x hx) (h0 x hx) -theorem SmoothWithinAt.inv₀ (hf : SmoothWithinAt I' I f s a) (ha : f a ≠ 0) : - SmoothWithinAt I' I (fun x => (f x)⁻¹) s a := - ContMDiffWithinAt.inv₀ hf ha - -theorem SmoothAt.inv₀ (hf : SmoothAt I' I f a) (ha : f a ≠ 0) : - SmoothAt I' I (fun x => (f x)⁻¹) a := - ContMDiffAt.inv₀ hf ha - -theorem Smooth.inv₀ (hf : Smooth I' I f) (h0 : ∀ x, f x ≠ 0) : Smooth I' I fun x => (f x)⁻¹ := - ContMDiff.inv₀ hf h0 - -theorem SmoothOn.inv₀ (hf : SmoothOn I' I f s) (h0 : ∀ x ∈ s, f x ≠ 0) : - SmoothOn I' I (fun x => (f x)⁻¹) s := - ContMDiffOn.inv₀ hf h0 +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.inv₀ := ContMDiffWithinAt.inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothAt.inv₀ := ContMDiffAt.inv₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn.inv₀ := ContMDiffOn.inv₀ +@[deprecated (since := "2024-11-21")] alias Smooth.inv₀ := ContMDiff.inv₀ end SmoothInv₀ @@ -313,20 +291,9 @@ theorem ContMDiffAt.div₀ (hf : ContMDiffAt I' I n f a) (hg : ContMDiffAt I' I theorem ContMDiff.div₀ (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) (h₀ : ∀ x, g x ≠ 0) : ContMDiff I' I n (f / g) := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv₀ h₀) -theorem SmoothWithinAt.div₀ (hf : SmoothWithinAt I' I f s a) - (hg : SmoothWithinAt I' I g s a) (h₀ : g a ≠ 0) : SmoothWithinAt I' I (f / g) s a := - ContMDiffWithinAt.div₀ hf hg h₀ - -theorem SmoothOn.div₀ (hf : SmoothOn I' I f s) (hg : SmoothOn I' I g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : - SmoothOn I' I (f / g) s := - ContMDiffOn.div₀ hf hg h₀ - -theorem SmoothAt.div₀ (hf : SmoothAt I' I f a) (hg : SmoothAt I' I g a) (h₀ : g a ≠ 0) : - SmoothAt I' I (f / g) a := - ContMDiffAt.div₀ hf hg h₀ - -theorem Smooth.div₀ (hf : Smooth I' I f) (hg : Smooth I' I g) (h₀ : ∀ x, g x ≠ 0) : - Smooth I' I (f / g) := - ContMDiff.div₀ hf hg h₀ +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div₀ := ContMDiffWithinAt.div₀ +@[deprecated (since := "2024-11-21")] alias SmoothAt.div₀ := ContMDiffAt.div₀ +@[deprecated (since := "2024-11-21")] alias SmoothOn.div₀ := ContMDiffOn.div₀ +@[deprecated (since := "2024-11-21")] alias Smooth.div₀ := ContMDiff.div₀ end Div diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 5439b39e461ed..3808ac40a51c1 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -45,7 +45,7 @@ class SmoothAdd {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topol {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Add G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothManifoldWithCorners I G : Prop where - smooth_add : Smooth (I.prod I) I fun p : G × G => p.1 + p.2 + smooth_add : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 + p.2 -- See note [Design choices about smooth algebraic structures] /-- Basic hypothesis to talk about a smooth (Lie) monoid or a smooth semigroup. @@ -56,7 +56,7 @@ class SmoothMul {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [Topol {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] (I : ModelWithCorners 𝕜 E H) (G : Type*) [Mul G] [TopologicalSpace G] [ChartedSpace H G] extends SmoothManifoldWithCorners I G : Prop where - smooth_mul : Smooth (I.prod I) I fun p : G × G => p.1 * p.2 + smooth_mul : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 * p.2 section SmoothMul @@ -71,16 +71,19 @@ section variable (I) @[to_additive] -theorem smooth_mul : Smooth (I.prod I) I fun p : G × G => p.1 * p.2 := +theorem contMDiff_mul : ContMDiff (I.prod I) I ⊤ fun p : G × G => p.1 * p.2 := SmoothMul.smooth_mul +@[deprecated (since := "2024-11-20")] alias smooth_mul := contMDiff_mul +@[deprecated (since := "2024-11-20")] alias smooth_add := contMDiff_add + include I in /-- If the multiplication is smooth, then it is continuous. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]. -/ @[to_additive "If the addition is smooth, then it is continuous. This is not an instance for technical reasons, see note [Design choices about smooth algebraic structures]."] theorem continuousMul_of_smooth : ContinuousMul G := - ⟨(smooth_mul I).continuous⟩ + ⟨(contMDiff_mul I).continuous⟩ end @@ -91,7 +94,7 @@ variable {f g : M → G} {s : Set M} {x : M} {n : ℕ∞} @[to_additive] theorem ContMDiffWithinAt.mul (hf : ContMDiffWithinAt I' I n f s x) (hg : ContMDiffWithinAt I' I n g s x) : ContMDiffWithinAt I' I n (f * g) s x := - ((smooth_mul I).smoothAt.of_le le_top).comp_contMDiffWithinAt x (hf.prod_mk hg) + ((contMDiff_mul I).contMDiffAt.of_le le_top).comp_contMDiffWithinAt x (hf.prod_mk hg) @[to_additive] nonrec theorem ContMDiffAt.mul (hf : ContMDiffAt I' I n f x) (hg : ContMDiffAt I' I n g x) : @@ -106,52 +109,51 @@ theorem ContMDiffOn.mul (hf : ContMDiffOn I' I n f s) (hg : ContMDiffOn I' I n g theorem ContMDiff.mul (hf : ContMDiff I' I n f) (hg : ContMDiff I' I n g) : ContMDiff I' I n (f * g) := fun x => (hf x).mul (hg x) -@[to_additive] -nonrec theorem SmoothWithinAt.mul (hf : SmoothWithinAt I' I f s x) - (hg : SmoothWithinAt I' I g s x) : SmoothWithinAt I' I (f * g) s x := - hf.mul hg - -@[to_additive] -nonrec theorem SmoothAt.mul (hf : SmoothAt I' I f x) (hg : SmoothAt I' I g x) : - SmoothAt I' I (f * g) x := - hf.mul hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.mul := ContMDiffWithinAt.mul +@[deprecated (since := "2024-11-21")] alias SmoothAt.mul := ContMDiffAt.mul +@[deprecated (since := "2024-11-21")] alias SmoothOn.mul := ContMDiffOn.mul +@[deprecated (since := "2024-11-21")] alias Smooth.mul := ContMDiff.mul -@[to_additive] -nonrec theorem SmoothOn.mul (hf : SmoothOn I' I f s) (hg : SmoothOn I' I g s) : - SmoothOn I' I (f * g) s := - hf.mul hg +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.add := ContMDiffWithinAt.add +@[deprecated (since := "2024-11-21")] alias SmoothAt.add := ContMDiffAt.add +@[deprecated (since := "2024-11-21")] alias SmoothOn.add := ContMDiffOn.add +@[deprecated (since := "2024-11-21")] alias Smooth.add := ContMDiff.add @[to_additive] -nonrec theorem Smooth.mul (hf : Smooth I' I f) (hg : Smooth I' I g) : Smooth I' I (f * g) := - hf.mul hg +theorem contMDiff_mul_left {a : G} : ContMDiff I I n (a * ·) := + contMDiff_const.mul contMDiff_id -@[to_additive] -theorem smooth_mul_left {a : G} : Smooth I I fun b : G => a * b := - smooth_const.mul smooth_id +@[deprecated (since := "2024-11-21")] alias smooth_mul_left := contMDiff_mul_left +@[deprecated (since := "2024-11-21")] alias smooth_add_left := contMDiff_add_left @[to_additive] -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 +@[to_additive] theorem mdifferentiable_mul_left {a : G} : MDifferentiable I I (a * ·) := contMDiff_mul_left.mdifferentiable le_rfl +@[to_additive] 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 +@[to_additive] +theorem contMDiff_mul_right {a : G} : ContMDiff I I n (· * a) := + contMDiff_id.mul contMDiff_const + +@[deprecated (since := "2024-11-21")] alias smooth_mul_right := contMDiff_mul_right +@[deprecated (since := "2024-11-21")] alias smooth_add_right := contMDiff_add_right +@[to_additive] theorem contMDiffAt_mul_right {a b : G} : ContMDiffAt I I n (· * a) b := contMDiff_mul_right.contMDiffAt +@[to_additive] theorem mdifferentiable_mul_right {a : G} : MDifferentiable I I (· * a) := contMDiff_mul_right.mdifferentiable le_rfl +@[to_additive] theorem mdifferentiableAt_mul_right {a b : G} : MDifferentiableAt I I (· * a) b := contMDiffAt_mul_right.mdifferentiableAt le_rfl @@ -163,13 +165,13 @@ variable (I) (g h : G) Lemmas involving `smoothLeftMul` with the notation `𝑳` usually use `L` instead of `𝑳` in the names. -/ def smoothLeftMul : C^∞⟮I, G; I, G⟯ := - ⟨leftMul g, smooth_mul_left⟩ + ⟨leftMul g, contMDiff_mul_left⟩ /-- Right multiplication by `g`. It is meant to mimic the usual notation in Lie groups. Lemmas involving `smoothRightMul` with the notation `𝑹` usually use `R` instead of `𝑹` in the names. -/ def smoothRightMul : C^∞⟮I, G; I, G⟯ := - ⟨rightMul g, smooth_mul_right⟩ + ⟨rightMul g, contMDiff_mul_right⟩ -- Left multiplication. The abbreviation is `MIL`. scoped[LieGroup] notation "𝑳" => smoothLeftMul @@ -222,8 +224,8 @@ instance SmoothMul.prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type* [ChartedSpace H' G'] [Mul G'] [SmoothMul I' G'] : SmoothMul (I.prod I') (G × G') := { SmoothManifoldWithCorners.prod G G' with smooth_mul := - ((smooth_fst.comp smooth_fst).smooth.mul (smooth_fst.comp smooth_snd)).prod_mk - ((smooth_snd.comp smooth_fst).smooth.mul (smooth_snd.comp smooth_snd)) } + ((contMDiff_fst.comp contMDiff_fst).mul (contMDiff_fst.comp contMDiff_snd)).prod_mk + ((contMDiff_snd.comp contMDiff_fst).mul (contMDiff_snd.comp contMDiff_snd)) } end SmoothMul @@ -236,16 +238,19 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS {G' : Type*} [Monoid G'] [TopologicalSpace G'] [ChartedSpace H' G'] [SmoothMul I' G'] @[to_additive] -theorem smooth_pow : ∀ n : ℕ, Smooth I I fun a : G => a ^ n - | 0 => by simp only [pow_zero]; exact smooth_const - | k + 1 => by simpa [pow_succ] using (smooth_pow _).mul smooth_id +theorem contMDiff_pow : ∀ n : ℕ, ContMDiff I I ⊤ fun a : G => a ^ n + | 0 => by simp only [pow_zero]; exact contMDiff_const + | k + 1 => by simpa [pow_succ] using (contMDiff_pow _).mul contMDiff_id + +@[deprecated (since := "2024-11-21")] alias smooth_pow := contMDiff_pow +@[deprecated (since := "2024-11-21")] alias smooth_nsmul := contMDiff_nsmul /-- Morphism of additive smooth monoids. -/ structure SmoothAddMonoidMorphism (I : ModelWithCorners 𝕜 E H) (I' : ModelWithCorners 𝕜 E' H') (G : Type*) [TopologicalSpace G] [ChartedSpace H G] [AddMonoid G] [SmoothAdd I G] (G' : Type*) [TopologicalSpace G'] [ChartedSpace H' G'] [AddMonoid G'] [SmoothAdd I' G'] extends G →+ G' where - smooth_toFun : Smooth I I' toFun + smooth_toFun : ContMDiff I I' ⊤ toFun /-- Morphism of smooth monoids. -/ @[to_additive] @@ -253,11 +258,11 @@ structure SmoothMonoidMorphism (I : ModelWithCorners 𝕜 E H) (I' : ModelWithCo (G : Type*) [TopologicalSpace G] [ChartedSpace H G] [Monoid G] [SmoothMul I G] (G' : Type*) [TopologicalSpace G'] [ChartedSpace H' G'] [Monoid G'] [SmoothMul I' G'] extends G →* G' where - smooth_toFun : Smooth I I' toFun + smooth_toFun : ContMDiff I I' ⊤ toFun @[to_additive] instance : One (SmoothMonoidMorphism I I' G G') := - ⟨{ smooth_toFun := smooth_const + ⟨{ smooth_toFun := contMDiff_const toMonoidHom := 1 }⟩ @[to_additive] @@ -389,61 +394,43 @@ theorem contMDiff_finprod_cond (hc : ∀ i, p i → ContMDiff I' I n (f i)) simp only [← finprod_subtype_eq_finprod_cond] exact contMDiff_finprod (fun i => hc i i.2) (hf.comp_injective Subtype.coe_injective) -@[to_additive] -theorem smoothAt_finprod - (lf : LocallyFinite fun i ↦ mulSupport <| f i) (h : ∀ i, SmoothAt I' I (f i) x₀) : - SmoothAt I' I (fun x ↦ ∏ᶠ i, f i x) x₀ := - contMDiffWithinAt_finprod lf h +@[deprecated (since := "2024-11-21")] alias smoothAt_finprod := contMDiffAt_finprod +@[deprecated (since := "2024-11-21")] alias smoothAt_finsum := contMDiffAt_finsum -@[to_additive] -theorem smoothWithinAt_finset_prod' (h : ∀ i ∈ t, SmoothWithinAt I' I (f i) s x) : - SmoothWithinAt I' I (∏ i ∈ t, f i) s x := - contMDiffWithinAt_finset_prod' h +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_prod' := contMDiffWithinAt_finset_prod' +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_sum' := contMDiffWithinAt_finset_sum' -@[to_additive] -theorem smoothWithinAt_finset_prod (h : ∀ i ∈ t, SmoothWithinAt I' I (f i) s x) : - SmoothWithinAt I' I (fun x => ∏ i ∈ t, f i x) s x := - contMDiffWithinAt_finset_prod h -@[to_additive] -theorem smoothAt_finset_prod' (h : ∀ i ∈ t, SmoothAt I' I (f i) x) : - SmoothAt I' I (∏ i ∈ t, f i) x := - contMDiffAt_finset_prod' h +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_prod := contMDiffWithinAt_finset_prod +@[deprecated (since := "2024-11-21")] +alias smoothWithinAt_finset_sum := contMDiffWithinAt_finset_sum -@[to_additive] -theorem smoothAt_finset_prod (h : ∀ i ∈ t, SmoothAt I' I (f i) x) : - SmoothAt I' I (fun x => ∏ i ∈ t, f i x) x := - contMDiffAt_finset_prod h +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_prod' := contMDiffAt_finset_prod' +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_sum' := contMDiffAt_finset_sum' -@[to_additive] -theorem smoothOn_finset_prod' (h : ∀ i ∈ t, SmoothOn I' I (f i) s) : - SmoothOn I' I (∏ i ∈ t, f i) s := - contMDiffOn_finset_prod' h +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_prod := contMDiffAt_finset_prod +@[deprecated (since := "2024-11-21")] alias smoothAt_finset_sum := contMDiffAt_finset_sum -@[to_additive] -theorem smoothOn_finset_prod (h : ∀ i ∈ t, SmoothOn I' I (f i) s) : - SmoothOn I' I (fun x => ∏ i ∈ t, f i x) s := - contMDiffOn_finset_prod h +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_prod' := contMDiffOn_finset_prod' +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_sum' := contMDiffOn_finset_sum' -@[to_additive] -theorem smooth_finset_prod' (h : ∀ i ∈ t, Smooth I' I (f i)) : Smooth I' I (∏ i ∈ t, f i) := - contMDiff_finset_prod' h +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_prod := contMDiffOn_finset_prod +@[deprecated (since := "2024-11-21")] alias smoothOn_finset_sum := contMDiffOn_finset_sum -@[to_additive] -theorem smooth_finset_prod (h : ∀ i ∈ t, Smooth I' I (f i)) : - Smooth I' I fun x => ∏ i ∈ t, f i x := - contMDiff_finset_prod h +@[deprecated (since := "2024-11-21")] alias smooth_finset_prod' := contMDiffOn_finset_prod' +@[deprecated (since := "2024-11-21")] alias smooth_finset_sum' := contMDiffOn_finset_sum' -@[to_additive] -theorem smooth_finprod (h : ∀ i, Smooth I' I (f i)) - (hfin : LocallyFinite fun i => mulSupport (f i)) : Smooth I' I fun x => ∏ᶠ i, f i x := - contMDiff_finprod h hfin +@[deprecated (since := "2024-11-21")] alias smooth_finset_prod := contMDiff_finset_prod +@[deprecated (since := "2024-11-21")] alias smooth_finset_sum := contMDiff_finset_sum -@[to_additive] -theorem smooth_finprod_cond (hc : ∀ i, p i → Smooth I' I (f i)) - (hf : LocallyFinite fun i => mulSupport (f i)) : - Smooth I' I fun x => ∏ᶠ (i) (_ : p i), f i x := - contMDiff_finprod_cond hc hf +@[deprecated (since := "2024-11-21")] alias smooth_finprod := contMDiff_finprod +@[deprecated (since := "2024-11-21")] alias smooth_finsum := contMDiff_finsum + +@[deprecated (since := "2024-11-21")] alias smooth_finprod_cond := contMDiff_finprod_cond +@[deprecated (since := "2024-11-21")] alias smooth_finsum_cond := contMDiff_finsum_cond end CommMonoid @@ -488,23 +475,9 @@ theorem ContMDiffOn.div_const (hf : ContMDiffOn I' I n f s) : theorem ContMDiff.div_const (hf : ContMDiff I' I n f) : ContMDiff I' I n (fun x ↦ f x / c) := fun x => (hf x).div_const c -@[to_additive] -nonrec theorem SmoothWithinAt.div_const (hf : SmoothWithinAt I' I f s x) : - SmoothWithinAt I' I (fun x ↦ f x / c) s x := - hf.div_const c - -@[to_additive] -nonrec theorem SmoothAt.div_const (hf : SmoothAt I' I f x) : - SmoothAt I' I (fun x ↦ f x / c) x := - hf.div_const c - -@[to_additive] -nonrec theorem SmoothOn.div_const (hf : SmoothOn I' I f s) : - SmoothOn I' I (fun x ↦ f x / c) s := - hf.div_const c - -@[to_additive] -nonrec theorem Smooth.div_const (hf : Smooth I' I f) : Smooth I' I (fun x ↦ f x / c) := - hf.div_const c +@[deprecated (since := "2024-11-21")] alias SmoothWithinAt.div_const := ContMDiffWithinAt.div_const +@[deprecated (since := "2024-11-21")] alias SmoothAt.div_const := ContMDiffAt.div_const +@[deprecated (since := "2024-11-21")] alias SmoothOn.div_const := ContMDiffOn.div_const +@[deprecated (since := "2024-11-21")] alias Smooth.div_const := ContMDiff.div_const end DivConst diff --git a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean index 002ad1cf560cf..e8adb09c03d94 100644 --- a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean +++ b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean @@ -30,7 +30,7 @@ namespace SmoothMap @[to_additive] protected instance instMul {G : Type*} [Mul G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] : Mul C^∞⟮I, N; I', G⟯ := - ⟨fun f g => ⟨f * g, f.smooth.mul g.smooth⟩⟩ + ⟨fun f g => ⟨f * g, f.contMDiff.mul g.contMDiff⟩⟩ @[to_additive (attr := simp)] theorem coe_mul {G : Type*} [Mul G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] @@ -54,12 +54,12 @@ theorem coe_one {G : Type*} [One G] [TopologicalSpace G] [ChartedSpace H' G] : instance instNSMul {G : Type*} [AddMonoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothAdd I' G] : SMul ℕ C^∞⟮I, N; I', G⟯ where - smul n f := ⟨n • (f : N → G), (smooth_nsmul n).comp f.smooth⟩ + smul n f := ⟨n • (f : N → G), (contMDiff_nsmul n).comp f.contMDiff⟩ @[to_additive existing] instance instPow {G : Type*} [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] : Pow C^∞⟮I, N; I', G⟯ ℕ where - pow f n := ⟨(f : N → G) ^ n, (smooth_pow n).comp f.smooth⟩ + pow f n := ⟨(f : N → G) ^ n, (contMDiff_pow n).comp f.contMDiff⟩ @[to_additive (attr := simp)] theorem coe_pow {G : Type*} [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] @@ -104,9 +104,9 @@ variable (I N) `C^∞⟮I, N; I'', G''⟯`."] def compLeftMonoidHom {G' : Type*} [Monoid G'] [TopologicalSpace G'] [ChartedSpace H' G'] [SmoothMul I' G'] {G'' : Type*} [Monoid G''] [TopologicalSpace G''] [ChartedSpace H'' G''] - [SmoothMul I'' G''] (φ : G' →* G'') (hφ : Smooth I' I'' φ) : + [SmoothMul I'' G''] (φ : G' →* G'') (hφ : ContMDiff I' I'' ⊤ φ) : C^∞⟮I, N; I', G'⟯ →* C^∞⟮I, N; I'', G''⟯ where - toFun f := ⟨φ ∘ f, fun x => (hφ.smooth _).comp x (f.contMDiff x)⟩ + toFun f := ⟨φ ∘ f, hφ.comp f.contMDiff⟩ map_one' := by ext; show φ 1 = 1; simp map_mul' f g := by ext x; show φ (f x * g x) = φ (f x) * φ (g x); simp @@ -119,7 +119,7 @@ variable (I') {N} homomorphism from `C^∞⟮I, V; I', G⟯` to `C^∞⟮I, U; I', G⟯`."] def restrictMonoidHom (G : Type*) [Monoid G] [TopologicalSpace G] [ChartedSpace H' G] [SmoothMul I' G] {U V : Opens N} (h : U ≤ V) : C^∞⟮I, V; I', G⟯ →* C^∞⟮I, U; I', G⟯ where - toFun f := ⟨f ∘ Set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩ + toFun f := ⟨f ∘ Set.inclusion h, f.contMDiff.comp (contMDiff_inclusion h)⟩ map_one' := rfl map_mul' _ _ := rfl @@ -134,9 +134,9 @@ instance commMonoid {G : Type*} [CommMonoid G] [TopologicalSpace G] [ChartedSpac instance group {G : Type*} [Group G] [TopologicalSpace G] [ChartedSpace H' G] [LieGroup I' G] : Group C^∞⟮I, N; I', G⟯ := { SmoothMap.monoid with - inv := fun f => ⟨fun x => (f x)⁻¹, f.smooth.inv⟩ + inv := fun f => ⟨fun x => (f x)⁻¹, f.contMDiff.inv⟩ inv_mul_cancel := fun a => by ext; exact inv_mul_cancel _ - div := fun f g => ⟨f / g, f.smooth.div g.smooth⟩ + div := fun f g => ⟨f / g, f.contMDiff.div g.contMDiff⟩ div_eq_mul_inv := fun f g => by ext; exact div_eq_mul_inv _ _ } @[to_additive (attr := simp)] @@ -189,11 +189,11 @@ variable (I N) 'left-composition-by-`φ`' ring homomorphism from `C^∞⟮I, N; I', R'⟯` to `C^∞⟮I, N; I'', R''⟯`. -/ def compLeftRingHom {R' : Type*} [Ring R'] [TopologicalSpace R'] [ChartedSpace H' R'] [SmoothRing I' R'] {R'' : Type*} [Ring R''] [TopologicalSpace R''] [ChartedSpace H'' R''] - [SmoothRing I'' R''] (φ : R' →+* R'') (hφ : Smooth I' I'' φ) : + [SmoothRing I'' R''] (φ : R' →+* R'') (hφ : ContMDiff I' I'' ⊤ φ) : C^∞⟮I, N; I', R'⟯ →+* C^∞⟮I, N; I'', R''⟯ := { SmoothMap.compLeftMonoidHom I N φ.toMonoidHom hφ, SmoothMap.compLeftAddMonoidHom I N φ.toAddMonoidHom hφ with - toFun := fun f => ⟨φ ∘ f, fun x => (hφ.smooth _).comp x (f.contMDiff x)⟩ } + toFun := fun f => ⟨φ ∘ f, hφ.comp f.contMDiff⟩ } variable (I') {N} @@ -202,7 +202,7 @@ variable (I') {N} def restrictRingHom (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H' R] [SmoothRing I' R] {U V : Opens N} (h : U ≤ V) : C^∞⟮I, V; I', R⟯ →+* C^∞⟮I, U; I', R⟯ := { SmoothMap.restrictMonoidHom I I' R h, SmoothMap.restrictAddMonoidHom I I' R h with - toFun := fun f => ⟨f ∘ Set.inclusion h, f.smooth.comp (smooth_inclusion h)⟩ } + toFun := fun f => ⟨f ∘ Set.inclusion h, f.contMDiff.comp (contMDiff_inclusion h)⟩ } variable {I I'} @@ -232,7 +232,7 @@ field `𝕜` inherit a vector space structure. instance instSMul {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] : SMul 𝕜 C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := - ⟨fun r f => ⟨r • ⇑f, smooth_const.smul f.smooth⟩⟩ + ⟨fun r f => ⟨r • ⇑f, contMDiff_const.smul f.contMDiff⟩⟩ @[simp] theorem coe_smul {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] (r : 𝕜) @@ -272,7 +272,7 @@ variable {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] [SmoothRing 𝓘(𝕜 /-- Smooth constant functions as a `RingHom`. -/ def C : 𝕜 →+* C^∞⟮I, N; 𝓘(𝕜, A), A⟯ where - toFun := fun c : 𝕜 => ⟨fun _ => (algebraMap 𝕜 A) c, smooth_const⟩ + toFun := fun c : 𝕜 => ⟨fun _ => (algebraMap 𝕜 A) c, contMDiff_const⟩ map_one' := by ext; exact (algebraMap 𝕜 A).map_one map_mul' c₁ c₂ := by ext; exact (algebraMap 𝕜 A).map_mul _ _ map_zero' := by ext; exact (algebraMap 𝕜 A).map_zero @@ -280,7 +280,7 @@ def C : 𝕜 →+* C^∞⟮I, N; 𝓘(𝕜, A), A⟯ where instance algebra : Algebra 𝕜 C^∞⟮I, N; 𝓘(𝕜, A), A⟯ := { --SmoothMap.semiring with -- Porting note: Commented this out. - smul := fun r f => ⟨r • f, smooth_const.smul f.smooth⟩ + smul := fun r f => ⟨r • f, contMDiff_const.smul f.contMDiff⟩ toRingHom := SmoothMap.C commutes' := fun c f => by ext x; exact Algebra.commutes' _ _ smul_def' := fun c f => by ext x; exact Algebra.smul_def' _ _ } @@ -309,7 +309,7 @@ is naturally a vector space over the ring of smooth functions from `N` to `𝕜` instance instSMul' {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] : SMul C^∞⟮I, N; 𝕜⟯ C^∞⟮I, N; 𝓘(𝕜, V), V⟯ := - ⟨fun f g => ⟨fun x => f x • g x, Smooth.smul f.2 g.2⟩⟩ + ⟨fun f g => ⟨fun x => f x • g x, ContMDiff.smul f.2 g.2⟩⟩ @[simp] theorem smul_comp' {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] (f : C^∞⟮I'', N'; 𝕜⟯) diff --git a/Mathlib/Geometry/Manifold/Algebra/Structures.lean b/Mathlib/Geometry/Manifold/Algebra/Structures.lean index 11c9d9b1a863a..da593ec373f61 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Structures.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Structures.lean @@ -24,7 +24,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {H : Type*} [TopologicalS If `R` is a ring, then negation is automatically smooth, as it is multiplication with `-1`. -/ class SmoothRing (I : ModelWithCorners 𝕜 E H) (R : Type*) [Semiring R] [TopologicalSpace R] [ChartedSpace H R] extends SmoothAdd I R : Prop where - smooth_mul : Smooth (I.prod I) I fun p : R × R => p.1 * p.2 + smooth_mul : ContMDiff (I.prod I) I ⊤ fun p : R × R => p.1 * p.2 -- see Note [lower instance priority] instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H) (R : Type*) @@ -36,8 +36,8 @@ instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H instance (priority := 100) SmoothRing.toLieAddGroup (I : ModelWithCorners 𝕜 E H) (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H R] [SmoothRing I R] : LieAddGroup I R where compatible := StructureGroupoid.compatible (contDiffGroupoid ⊤ I) - smooth_add := smooth_add I - smooth_neg := by simpa only [neg_one_mul] using @smooth_mul_left 𝕜 _ H _ E _ _ I R _ _ _ _ (-1) + smooth_add := contMDiff_add I + smooth_neg := by simpa only [neg_one_mul] using contMDiff_mul_left (G := R) (a := -1) end SmoothRing @@ -46,7 +46,7 @@ instance (priority := 100) fieldSmoothRing {𝕜 : Type*} [NontriviallyNormedFie SmoothRing 𝓘(𝕜) 𝕜 := { normedSpaceLieAddGroup with smooth_mul := by - rw [smooth_iff] + rw [contMDiff_iff] refine ⟨continuous_mul, fun x y => ?_⟩ simp only [mfld_simps] rw [contDiffOn_univ] diff --git a/Mathlib/Geometry/Manifold/BumpFunction.lean b/Mathlib/Geometry/Manifold/BumpFunction.lean index fa6698b53773a..0c37faf328e8a 100644 --- a/Mathlib/Geometry/Manifold/BumpFunction.lean +++ b/Mathlib/Geometry/Manifold/BumpFunction.lean @@ -286,23 +286,28 @@ theorem nhds_basis_support {s : Set M} (hs : s ∈ 𝓝 c) : variable [SmoothManifoldWithCorners I M] /-- A smooth bump function is infinitely smooth. -/ -protected theorem smooth : Smooth I 𝓘(ℝ) f := by +protected theorem contMDiff : ContMDiff I 𝓘(ℝ) ⊤ f := by refine contMDiff_of_tsupport fun x hx => ?_ have : x ∈ (chartAt H c).source := f.tsupport_subset_chartAt_source hx refine ContMDiffAt.congr_of_eventuallyEq ?_ <| f.eqOn_source.eventuallyEq_of_mem <| (chartAt H c).open_source.mem_nhds this exact f.contDiffAt.contMDiffAt.comp _ (contMDiffAt_extChartAt' this) -protected theorem smoothAt {x} : SmoothAt I 𝓘(ℝ) f x := - f.smooth.smoothAt +@[deprecated (since := "2024-11-20")] alias smooth := SmoothBumpFunction.contMDiff + +protected theorem contMDiffAt {x} : ContMDiffAt I 𝓘(ℝ) ⊤ f x := + f.contMDiff.contMDiffAt + +@[deprecated (since := "2024-11-20")] alias smoothAt := SmoothBumpFunction.contMDiffAt protected theorem continuous : Continuous f := - f.smooth.continuous + f.contMDiff.continuous /-- If `f : SmoothBumpFunction I c` is a smooth bump function and `g : M → G` is a function smooth on the source of the chart at `c`, then `f • g` is smooth on the whole manifold. -/ -theorem smooth_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} - (hg : SmoothOn I 𝓘(ℝ, G) g (chartAt H c).source) : Smooth I 𝓘(ℝ, G) fun x => f x • g x := by +theorem contMDiff_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} + (hg : ContMDiffOn I 𝓘(ℝ, G) ⊤ g (chartAt H c).source) : + ContMDiff I 𝓘(ℝ, G) ⊤ fun x => f x • g x := by refine contMDiff_of_tsupport fun x hx => ?_ have : x ∈ (chartAt H c).source := -- Porting note: was a more readable `calc` @@ -311,6 +316,8 @@ theorem smooth_smul {G} [NormedAddCommGroup G] [NormedSpace ℝ G] {g : M → G} -- _ ⊆ tsupport f := tsupport_smul_subset_left _ _ -- _ ⊆ (chart_at _ c).source := f.tsupport_subset_chartAt_source f.tsupport_subset_chartAt_source <| tsupport_smul_subset_left _ _ hx - exact f.smoothAt.smul ((hg _ this).contMDiffAt <| (chartAt _ _).open_source.mem_nhds this) + exact f.contMDiffAt.smul ((hg _ this).contMDiffAt <| (chartAt _ _).open_source.mem_nhds this) + +@[deprecated (since := "2024-11-20")] alias smooth_smul := contMDiff_smul end SmoothBumpFunction diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 5f68ad7682a06..2036cda540d21 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -125,7 +125,7 @@ variable [ChartedSpace H M'] [IsM' : SmoothManifoldWithCorners I M'] theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M M'} (hf : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source) : - SmoothOn I I f f.source := by + ContMDiffOn I I ⊤ f f.source := by -- It suffices to show smoothness near each `x` apply contMDiffOn_of_locally_contMDiffOn intro x hx @@ -171,7 +171,7 @@ is a local structomorphism for `I`, if and only if it is manifold-smooth on the in both directions. -/ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source ↔ - SmoothOn I I f f.source ∧ SmoothOn I I f.symm f.target := by + ContMDiffOn I I ⊤ f f.source ∧ ContMDiffOn I I ⊤ f.symm f.target := by constructor · intro h refine ⟨isLocalStructomorphOn_contDiffGroupoid_iff_aux h, diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index 0b2a089aee1e6..f5506bc51360d 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -82,31 +82,21 @@ theorem ContMDiffWithinAt.comp_of_eq {t : Set M'} {g : M' → M''} {x : M} {y : (st : MapsTo f s t) (hx : f x = y) : ContMDiffWithinAt I I'' n (g ∘ f) s x := by subst hx; exact hg.comp x hf st -/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ -nonrec theorem SmoothWithinAt.comp {t : Set M'} {g : M' → M''} (x : M) - (hg : SmoothWithinAt I' I'' g t (f x)) (hf : SmoothWithinAt I I' f s x) (st : MapsTo f s t) : - SmoothWithinAt I I'' (g ∘ f) s x := - hg.comp x hf st +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.comp := ContMDiffWithinAt.comp /-- The composition of `C^n` functions on domains is `C^n`. -/ theorem ContMDiffOn.comp {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiffOn I I' n f s) (st : s ⊆ f ⁻¹' t) : ContMDiffOn I I'' n (g ∘ f) s := fun x hx => (hg _ (st hx)).comp x (hf x hx) st -/-- The composition of `C^∞` functions on domains is `C^∞`. -/ -nonrec theorem SmoothOn.comp {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : SmoothOn I I' f s) (st : s ⊆ f ⁻¹' t) : SmoothOn I I'' (g ∘ f) s := - hg.comp hf st +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp := ContMDiffOn.comp /-- The composition of `C^n` functions on domains is `C^n`. -/ theorem ContMDiffOn.comp' {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiffOn I I' n f s) : ContMDiffOn I I'' n (g ∘ f) (s ∩ f ⁻¹' t) := hg.comp (hf.mono inter_subset_left) inter_subset_right -/-- The composition of `C^∞` functions is `C^∞`. -/ -nonrec theorem SmoothOn.comp' {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : SmoothOn I I' f s) : SmoothOn I I'' (g ∘ f) (s ∩ f ⁻¹' t) := - hg.comp' hf +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp' := ContMDiffOn.comp' /-- The composition of `C^n` functions is `C^n`. -/ theorem ContMDiff.comp {g : M' → M''} (hg : ContMDiff I' I'' n g) (hf : ContMDiff I I' n f) : @@ -114,10 +104,7 @@ theorem ContMDiff.comp {g : M' → M''} (hg : ContMDiff I' I'' n g) (hf : ContMD rw [← contMDiffOn_univ] at hf hg ⊢ exact hg.comp hf subset_preimage_univ -/-- The composition of `C^∞` functions is `C^∞`. -/ -nonrec theorem Smooth.comp {g : M' → M''} (hg : Smooth I' I'' g) (hf : Smooth I I' f) : - Smooth I I'' (g ∘ f) := - hg.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.comp := ContMDiff.comp /-- The composition of `C^n` functions within domains at points is `C^n`. -/ theorem ContMDiffWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) @@ -125,11 +112,7 @@ theorem ContMDiffWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) ContMDiffWithinAt I I'' n (g ∘ f) (s ∩ f ⁻¹' t) x := hg.comp x (hf.mono inter_subset_left) inter_subset_right -/-- The composition of `C^∞` functions within domains at points is `C^∞`. -/ -nonrec theorem SmoothWithinAt.comp' {t : Set M'} {g : M' → M''} (x : M) - (hg : SmoothWithinAt I' I'' g t (f x)) (hf : SmoothWithinAt I I' f s x) : - SmoothWithinAt I I'' (g ∘ f) (s ∩ f ⁻¹' t) x := - hg.comp' x hf +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.comp' := ContMDiffWithinAt.comp' /-- `g ∘ f` is `C^n` within `s` at `x` if `g` is `C^n` at `f x` and `f` is `C^n` within `s` at `x`. -/ @@ -138,11 +121,8 @@ theorem ContMDiffAt.comp_contMDiffWithinAt {g : M' → M''} (x : M) ContMDiffWithinAt I I'' n (g ∘ f) s x := hg.comp x hf (mapsTo_univ _ _) -/-- `g ∘ f` is `C^∞` within `s` at `x` if `g` is `C^∞` at `f x` and -`f` is `C^∞` within `s` at `x`. -/ -theorem SmoothAt.comp_smoothWithinAt {g : M' → M''} (x : M) (hg : SmoothAt I' I'' g (f x)) - (hf : SmoothWithinAt I I' f s x) : SmoothWithinAt I I'' (g ∘ f) s x := - hg.comp_contMDiffWithinAt x hf +@[deprecated (since := "2024-11-20")] +alias SmoothAt.comp_smoothWithinAt := ContMDiffAt.comp_contMDiffWithinAt /-- The composition of `C^n` functions at points is `C^n`. -/ nonrec theorem ContMDiffAt.comp {g : M' → M''} (x : M) (hg : ContMDiffAt I' I'' n g (f x)) @@ -154,26 +134,19 @@ theorem ContMDiffAt.comp_of_eq {g : M' → M''} {x : M} {y : M'} (hg : ContMDiff (hf : ContMDiffAt I I' n f x) (hx : f x = y) : ContMDiffAt I I'' n (g ∘ f) x := by subst hx; exact hg.comp x hf -/-- The composition of `C^∞` functions at points is `C^∞`. -/ -nonrec theorem SmoothAt.comp {g : M' → M''} (x : M) (hg : SmoothAt I' I'' g (f x)) - (hf : SmoothAt I I' f x) : SmoothAt I I'' (g ∘ f) x := - hg.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.comp := ContMDiffAt.comp theorem ContMDiff.comp_contMDiffOn {f : M → M'} {g : M' → M''} {s : Set M} (hg : ContMDiff I' I'' n g) (hf : ContMDiffOn I I' n f s) : ContMDiffOn I I'' n (g ∘ f) s := hg.contMDiffOn.comp hf Set.subset_preimage_univ -theorem Smooth.comp_smoothOn {f : M → M'} {g : M' → M''} {s : Set M} (hg : Smooth I' I'' g) - (hf : SmoothOn I I' f s) : SmoothOn I I'' (g ∘ f) s := - hg.smoothOn.comp hf Set.subset_preimage_univ +@[deprecated (since := "2024-11-20")] alias Smooth.comp_smoothOn := ContMDiff.comp_contMDiffOn theorem ContMDiffOn.comp_contMDiff {t : Set M'} {g : M' → M''} (hg : ContMDiffOn I' I'' n g t) (hf : ContMDiff I I' n f) (ht : ∀ x, f x ∈ t) : ContMDiff I I'' n (g ∘ f) := contMDiffOn_univ.mp <| hg.comp hf.contMDiffOn fun x _ => ht x -theorem SmoothOn.comp_smooth {t : Set M'} {g : M' → M''} (hg : SmoothOn I' I'' g t) - (hf : Smooth I I' f) (ht : ∀ x, f x ∈ t) : Smooth I I'' (g ∘ f) := - hg.comp_contMDiff hf ht +@[deprecated (since := "2024-11-20")] alias SmoothOn.comp_smooth := ContMDiffOn.comp_contMDiff end Composition @@ -185,26 +158,22 @@ theorem contMDiff_id : ContMDiff I I n (id : M → M) := ContMDiff.of_le ((contDiffWithinAt_localInvariantProp ∞).liftProp_id contDiffWithinAtProp_id) le_top -theorem smooth_id : Smooth I I (id : M → M) := - contMDiff_id +@[deprecated (since := "2024-11-20")] alias smooth_id := contMDiff_id theorem contMDiffOn_id : ContMDiffOn I I n (id : M → M) s := contMDiff_id.contMDiffOn -theorem smoothOn_id : SmoothOn I I (id : M → M) s := - contMDiffOn_id +@[deprecated (since := "2024-11-20")] alias smoothOn_id := contMDiffOn_id theorem contMDiffAt_id : ContMDiffAt I I n (id : M → M) x := contMDiff_id.contMDiffAt -theorem smoothAt_id : SmoothAt I I (id : M → M) x := - contMDiffAt_id +@[deprecated (since := "2024-11-20")] alias smoothAt_id := contMDiffAt_id theorem contMDiffWithinAt_id : ContMDiffWithinAt I I n (id : M → M) s x := contMDiffAt_id.contMDiffWithinAt -theorem smoothWithinAt_id : SmoothWithinAt I I (id : M → M) s x := - contMDiffWithinAt_id +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_id := contMDiffWithinAt_id end id @@ -223,11 +192,10 @@ theorem contMDiff_const : ContMDiff I I' n fun _ : M => c := by theorem contMDiff_one [One M'] : ContMDiff I I' n (1 : M → M') := by simp only [Pi.one_def, contMDiff_const] -theorem smooth_const : Smooth I I' fun _ : M => c := - contMDiff_const +@[deprecated (since := "2024-11-20")] alias smooth_const := contMDiff_const -@[to_additive] -theorem smooth_one [One M'] : Smooth I I' (1 : M → M') := by simp only [Pi.one_def, smooth_const] +@[deprecated (since := "2024-11-20")] alias smooth_one := contMDiff_one +@[deprecated (since := "2024-11-20")] alias smooth_zero := contMDiff_zero theorem contMDiffOn_const : ContMDiffOn I I' n (fun _ : M => c) s := contMDiff_const.contMDiffOn @@ -236,12 +204,11 @@ theorem contMDiffOn_const : ContMDiffOn I I' n (fun _ : M => c) s := theorem contMDiffOn_one [One M'] : ContMDiffOn I I' n (1 : M → M') s := contMDiff_one.contMDiffOn -theorem smoothOn_const : SmoothOn I I' (fun _ : M => c) s := - contMDiffOn_const +@[deprecated (since := "2024-11-20")] alias smoothOn_const := contMDiffOn_const + +@[deprecated (since := "2024-11-20")] alias smoothOn_one := contMDiffOn_one +@[deprecated (since := "2024-11-20")] alias smoothOn_zero := contMDiffOn_zero -@[to_additive] -theorem smoothOn_one [One M'] : SmoothOn I I' (1 : M → M') s := - contMDiffOn_one theorem contMDiffAt_const : ContMDiffAt I I' n (fun _ : M => c) x := contMDiff_const.contMDiffAt @@ -250,12 +217,10 @@ theorem contMDiffAt_const : ContMDiffAt I I' n (fun _ : M => c) x := theorem contMDiffAt_one [One M'] : ContMDiffAt I I' n (1 : M → M') x := contMDiff_one.contMDiffAt -theorem smoothAt_const : SmoothAt I I' (fun _ : M => c) x := - contMDiffAt_const +@[deprecated (since := "2024-11-20")] alias smoothAt_const := contMDiffAt_const -@[to_additive] -theorem smoothAt_one [One M'] : SmoothAt I I' (1 : M → M') x := - contMDiffAt_one +@[deprecated (since := "2024-11-20")] alias smoothAt_one := contMDiffAt_one +@[deprecated (since := "2024-11-20")] alias smoothAt_zero := contMDiffAt_zero theorem contMDiffWithinAt_const : ContMDiffWithinAt I I' n (fun _ : M => c) s x := contMDiffAt_const.contMDiffWithinAt @@ -264,12 +229,10 @@ theorem contMDiffWithinAt_const : ContMDiffWithinAt I I' n (fun _ : M => c) s x theorem contMDiffWithinAt_one [One M'] : ContMDiffWithinAt I I' n (1 : M → M') s x := contMDiffAt_const.contMDiffWithinAt -theorem smoothWithinAt_const : SmoothWithinAt I I' (fun _ : M => c) s x := - contMDiffWithinAt_const +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_const := contMDiffWithinAt_const -@[to_additive] -theorem smoothWithinAt_one [One M'] : SmoothWithinAt I I' (1 : M → M') s x := - contMDiffWithinAt_one +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_one := contMDiffWithinAt_one +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_zero := contMDiffWithinAt_zero end id @@ -305,12 +268,14 @@ section Inclusion open TopologicalSpace -theorem contMdiffAt_subtype_iff {n : ℕ∞} {U : Opens M} {f : M → M'} {x : U} : +theorem contMDiffAt_subtype_iff {n : ℕ∞} {U : Opens M} {f : M → M'} {x : U} : ContMDiffAt I I' n (fun x : U ↦ f x) x ↔ ContMDiffAt I I' n f x := ((contDiffWithinAt_localInvariantProp n).liftPropAt_iff_comp_subtype_val _ _).symm +@[deprecated (since := "2024-11-20")] alias contMdiffAt_subtype_iff := contMDiffAt_subtype_iff + theorem contMDiff_subtype_val {n : ℕ∞} {U : Opens M} : ContMDiff I I n (Subtype.val : U → M) := - fun _ ↦ contMdiffAt_subtype_iff.mpr contMDiffAt_id + fun _ ↦ contMDiffAt_subtype_iff.mpr contMDiffAt_id @[to_additive] theorem ContMDiff.extend_one [T2Space M] [One M'] {n : ℕ∞} {U : Opens M} {f : U → M'} @@ -319,7 +284,7 @@ theorem ContMDiff.extend_one [T2Space M] [One M'] {n : ℕ∞} {U : Opens M} {f refine contMDiff_of_mulTSupport (fun x h ↦ ?_) _ lift x to U using Subtype.coe_image_subset _ _ (supp.mulTSupport_extend_one_subset continuous_subtype_val h) - rw [← contMdiffAt_subtype_iff] + rw [← contMDiffAt_subtype_iff] simp_rw [← comp_def] rw [extend_comp Subtype.val_injective] exact diff.contMDiffAt @@ -333,19 +298,14 @@ theorem contMDiff_inclusion {n : ℕ∞} {U V : Opens M} (h : U ≤ V) : rw [Set.univ_inter] exact contDiffWithinAt_id.congr I.rightInvOn (congr_arg I (I.left_inv y)) -theorem smooth_subtype_iff {U : Opens M} {f : M → M'} {x : U} : - SmoothAt I I' (fun x : U ↦ f x) x ↔ SmoothAt I I' f x := contMdiffAt_subtype_iff +@[deprecated (since := "2024-11-20")] alias smooth_subtype_iff := contMDiffAt_subtype_iff -theorem smooth_subtype_val {U : Opens M} : Smooth I I (Subtype.val : U → M) := contMDiff_subtype_val +@[deprecated (since := "2024-11-20")] alias smooth_subtype_val := contMDiff_subtype_val -@[to_additive] -theorem Smooth.extend_one [T2Space M] [One M'] {U : Opens M} {f : U → M'} - (supp : HasCompactMulSupport f) (diff : Smooth I I' f) : Smooth I I' (Subtype.val.extend f 1) := - ContMDiff.extend_one supp diff +@[deprecated (since := "2024-11-20")] alias Smooth.extend_one := ContMDiff.extend_one +@[deprecated (since := "2024-11-20")] alias Smooth.extend_zero := ContMDiff.extend_zero -theorem smooth_inclusion {U V : Opens M} (h : U ≤ V) : - Smooth I I (Opens.inclusion h : U → V) := - contMDiff_inclusion h +@[deprecated (since := "2024-11-20")] alias smooth_inclusion := contMDiff_inclusion end Inclusion diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index f0917e558dc3c..aa5e8d957cfd7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -155,11 +155,7 @@ read in the preferred chart at this point. -/ def ContMDiffWithinAt (n : ℕ∞) (f : M → M') (s : Set M) (x : M) := LiftPropWithinAt (ContDiffWithinAtProp I I' n) f s x -variable (I I') in -/-- Abbreviation for `ContMDiffWithinAt I I' ⊤ f s x`. See also documentation for `Smooth`. --/ -abbrev SmoothWithinAt (f : M → M') (s : Set M) (x : M) := - ContMDiffWithinAt I I' ⊤ f s x +@[deprecated (since := "024-11-21")] alias SmoothWithinAt := ContMDiffWithinAt variable (I I') in /-- A function is `n` times continuously differentiable at a point in a manifold if @@ -175,10 +171,7 @@ theorem contMDiffAt_iff {n : ℕ∞} {f : M → M'} {x : M} : (extChartAt I x x) := liftPropAt_iff.trans <| by rw [ContDiffWithinAtProp, preimage_univ, univ_inter]; rfl -variable (I I') in -/-- Abbreviation for `ContMDiffAt I I' ⊤ f x`. See also documentation for `Smooth`. -/ -abbrev SmoothAt (f : M → M') (x : M) := - ContMDiffAt I I' ⊤ f x +@[deprecated (since := "024-11-21")] alias SmoothAt := ContMDiffAt variable (I I') in /-- A function is `n` times continuously differentiable in a set of a manifold if it is continuous @@ -187,10 +180,7 @@ around these points. -/ def ContMDiffOn (n : ℕ∞) (f : M → M') (s : Set M) := ∀ x ∈ s, ContMDiffWithinAt I I' n f s x -variable (I I') in -/-- Abbreviation for `ContMDiffOn I I' ⊤ f s`. See also documentation for `Smooth`. -/ -abbrev SmoothOn (f : M → M') (s : Set M) := - ContMDiffOn I I' ⊤ f s +@[deprecated (since := "024-11-21")] alias SmoothOn := ContMDiffOn variable (I I') in /-- A function is `n` times continuously differentiable in a manifold if it is continuous @@ -199,16 +189,8 @@ around these points. -/ def ContMDiff (n : ℕ∞) (f : M → M') := ∀ x, ContMDiffAt I I' n f x -variable (I I') in -/-- Abbreviation for `ContMDiff I I' ⊤ f`. -Short note to work with these abbreviations: a lemma of the form `ContMDiffFoo.bar` will -apply fine to an assumption `SmoothFoo` using dot notation or normal notation. -If the consequence `bar` of the lemma involves `ContDiff`, it is still better to restate -the lemma replacing `ContDiff` with `Smooth` both in the assumption and in the conclusion, -to make it possible to use `Smooth` consistently. -This also applies to `SmoothAt`, `SmoothOn` and `SmoothWithinAt`. -/ -abbrev Smooth (f : M → M') := - ContMDiff I I' ⊤ f +@[deprecated (since := "024-11-21")] alias Smooth := ContMDiff + /-! ### Deducing smoothness from higher smoothness -/ @@ -228,49 +210,38 @@ theorem ContMDiff.of_le (hf : ContMDiff I I' n f) (le : m ≤ n) : ContMDiff I I /-! ### Basic properties of smooth functions between manifolds -/ -theorem ContMDiff.smooth (h : ContMDiff I I' ⊤ f) : Smooth I I' f := - h +@[deprecated (since := "2024-11-20")] alias ContMDiff.smooth := ContMDiff.of_le -theorem Smooth.contMDiff (h : Smooth I I' f) : ContMDiff I I' n f := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias Smooth.contMDiff := ContMDiff.of_le -theorem ContMDiffOn.smoothOn (h : ContMDiffOn I I' ⊤ f s) : SmoothOn I I' f s := - h +@[deprecated (since := "2024-11-20")] alias ContMDiffOn.smoothOn := ContMDiffOn.of_le -theorem SmoothOn.contMDiffOn (h : SmoothOn I I' f s) : ContMDiffOn I I' n f s := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias SmoothOn.contMDiffOn := ContMDiffOn.of_le -theorem ContMDiffAt.smoothAt (h : ContMDiffAt I I' ⊤ f x) : SmoothAt I I' f x := - h +@[deprecated (since := "2024-11-20")] alias ContMDiffAt.smoothAt := ContMDiffAt.of_le -theorem SmoothAt.contMDiffAt (h : SmoothAt I I' f x) : ContMDiffAt I I' n f x := - h.of_le le_top +@[deprecated (since := "2024-11-20")] alias SmoothAt.contMDiffAt := ContMDiffOn.of_le -theorem ContMDiffWithinAt.smoothWithinAt (h : ContMDiffWithinAt I I' ⊤ f s x) : - SmoothWithinAt I I' f s x := - h +@[deprecated (since := "2024-11-20")] +alias ContMDiffWithinAt.smoothWithinAt := ContMDiffWithinAt.of_le -theorem SmoothWithinAt.contMDiffWithinAt (h : SmoothWithinAt I I' f s x) : - ContMDiffWithinAt I I' n f s x := - h.of_le le_top +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.contMDiffWithinAt := ContMDiffWithinAt.of_le theorem ContMDiff.contMDiffAt (h : ContMDiff I I' n f) : ContMDiffAt I I' n f x := h x -theorem Smooth.smoothAt (h : Smooth I I' f) : SmoothAt I I' f x := - ContMDiff.contMDiffAt h +@[deprecated (since := "2024-11-20")] alias Smooth.smoothAt := ContMDiff.contMDiffAt theorem contMDiffWithinAt_univ : ContMDiffWithinAt I I' n f univ x ↔ ContMDiffAt I I' n f x := Iff.rfl -theorem smoothWithinAt_univ : SmoothWithinAt I I' f univ x ↔ SmoothAt I I' f x := - contMDiffWithinAt_univ +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_univ := contMDiffWithinAt_univ theorem contMDiffOn_univ : ContMDiffOn I I' n f univ ↔ ContMDiff I I' n f := by simp only [ContMDiffOn, ContMDiff, contMDiffWithinAt_univ, forall_prop_of_true, mem_univ] -theorem smoothOn_univ : SmoothOn I I' f univ ↔ Smooth I I' f := - contMDiffOn_univ +@[deprecated (since := "2024-11-20")] alias smoothOn_univ := contMDiffOn_univ /-- One can reformulate smoothness within a set at a point as continuity within this set at this point, and smoothness in the corresponding extended chart. -/ @@ -315,26 +286,18 @@ theorem contMDiffWithinAt_iff_target : chartAt_self_eq, PartialHomeomorph.refl_apply, id_comp] rfl -theorem smoothWithinAt_iff : - SmoothWithinAt I I' f s x ↔ - ContinuousWithinAt f s x ∧ - ContDiffWithinAt 𝕜 ∞ (extChartAt I' (f x) ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).symm ⁻¹' s ∩ range I) (extChartAt I x x) := - contMDiffWithinAt_iff +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_iff := contMDiffWithinAt_iff -theorem smoothWithinAt_iff_target : - SmoothWithinAt I I' f s x ↔ - ContinuousWithinAt f s x ∧ SmoothWithinAt I 𝓘(𝕜, E') (extChartAt I' (f x) ∘ f) s x := - contMDiffWithinAt_iff_target +@[deprecated (since := "2024-11-20")] +alias smoothWithinAt_iff_target := contMDiffWithinAt_iff_target theorem contMDiffAt_iff_target {x : M} : ContMDiffAt I I' n f x ↔ ContinuousAt f x ∧ ContMDiffAt I 𝓘(𝕜, E') n (extChartAt I' (f x) ∘ f) x := by rw [ContMDiffAt, ContMDiffAt, contMDiffWithinAt_iff_target, continuousWithinAt_univ] -theorem smoothAt_iff_target {x : M} : - SmoothAt I I' f x ↔ ContinuousAt f x ∧ SmoothAt I 𝓘(𝕜, E') (extChartAt I' (f x) ∘ f) x := - contMDiffAt_iff_target +@[deprecated (since := "2024-11-20")] alias smoothAt_iff_target := contMDiffAt_iff_target + section SmoothManifoldWithCorners @@ -533,20 +496,10 @@ theorem contMDiffOn_iff_target : simp · exact fun h' x y => (h' y).2 x 0 -theorem smoothOn_iff : - SmoothOn I I' f s ↔ - ContinuousOn f s ∧ - ∀ (x : M) (y : M'), - ContDiffOn 𝕜 ⊤ (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).target ∩ - (extChartAt I x).symm ⁻¹' (s ∩ f ⁻¹' (extChartAt I' y).source)) := - contMDiffOn_iff +@[deprecated (since := "2024-11-20")] alias smoothOn_iff := contMDiffOn_iff + +@[deprecated (since := "2024-11-20")] alias smoothOn_iff_target := contMDiffOn_iff_target -theorem smoothOn_iff_target : - SmoothOn I I' f s ↔ - ContinuousOn f s ∧ - ∀ y : M', SmoothOn I 𝓘(𝕜, E') (extChartAt I' y ∘ f) (s ∩ f ⁻¹' (extChartAt I' y).source) := - contMDiffOn_iff_target /-- One can reformulate smoothness as continuity and smoothness in any extended chart. -/ theorem contMDiff_iff : @@ -567,20 +520,9 @@ theorem contMDiff_iff_target : rw [← contMDiffOn_univ, contMDiffOn_iff_target] simp [continuous_iff_continuousOn_univ] -theorem smooth_iff : - Smooth I I' f ↔ - Continuous f ∧ - ∀ (x : M) (y : M'), - ContDiffOn 𝕜 ⊤ (extChartAt I' y ∘ f ∘ (extChartAt I x).symm) - ((extChartAt I x).target ∩ - (extChartAt I x).symm ⁻¹' (f ⁻¹' (extChartAt I' y).source)) := - contMDiff_iff +@[deprecated (since := "2024-11-20")] alias smooth_iff := contMDiff_iff -theorem smooth_iff_target : - Smooth I I' f ↔ - Continuous f ∧ - ∀ y : M', SmoothOn I 𝓘(𝕜, E') (extChartAt I' y ∘ f) (f ⁻¹' (extChartAt I' y).source) := - contMDiff_iff_target +@[deprecated (since := "2024-11-20")] alias smooth_iff_target := contMDiff_iff_target end SmoothManifoldWithCorners @@ -619,17 +561,17 @@ theorem ContMDiff.continuous (hf : ContMDiff I I' n f) : Continuous f := /-! ### `C^∞` smoothness -/ theorem contMDiffWithinAt_top : - SmoothWithinAt I I' f s x ↔ ∀ n : ℕ, ContMDiffWithinAt I I' n f s x := + ContMDiffWithinAt I I' ⊤ f s x ↔ ∀ n : ℕ, ContMDiffWithinAt I I' n f s x := ⟨fun h n => ⟨h.1, contDiffWithinAt_top.1 h.2 n⟩, fun H => ⟨(H 0).1, contDiffWithinAt_top.2 fun n => (H n).2⟩⟩ -theorem contMDiffAt_top : SmoothAt I I' f x ↔ ∀ n : ℕ, ContMDiffAt I I' n f x := +theorem contMDiffAt_top : ContMDiffAt I I' ⊤ f x ↔ ∀ n : ℕ, ContMDiffAt I I' n f x := contMDiffWithinAt_top -theorem contMDiffOn_top : SmoothOn I I' f s ↔ ∀ n : ℕ, ContMDiffOn I I' n f s := +theorem contMDiffOn_top : ContMDiffOn I I' ⊤ f s ↔ ∀ n : ℕ, ContMDiffOn I I' n f s := ⟨fun h _ => h.of_le le_top, fun h x hx => contMDiffWithinAt_top.2 fun n => h n x hx⟩ -theorem contMDiff_top : Smooth I I' f ↔ ∀ n : ℕ, ContMDiff I I' n f := +theorem contMDiff_top : ContMDiff I I' ⊤ f ↔ ∀ n : ℕ, ContMDiff I I' n f := ⟨fun h _ => h.of_le le_top, fun h x => contMDiffWithinAt_top.2 fun n => h n x⟩ theorem contMDiffWithinAt_iff_nat : @@ -691,8 +633,7 @@ protected theorem ContMDiffAt.contMDiffWithinAt (hf : ContMDiffAt I I' n f x) : ContMDiffWithinAt I I' n f s x := ContMDiffWithinAt.mono hf (subset_univ _) -protected theorem SmoothAt.smoothWithinAt (hf : SmoothAt I I' f x) : SmoothWithinAt I I' f s x := - ContMDiffAt.contMDiffWithinAt hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.smoothWithinAt := ContMDiffAt.contMDiffWithinAt theorem ContMDiffOn.mono (hf : ContMDiffOn I I' n f s) (hts : t ⊆ s) : ContMDiffOn I I' n f t := fun x hx => (hf x (hts hx)).mono hts @@ -700,8 +641,7 @@ theorem ContMDiffOn.mono (hf : ContMDiffOn I I' n f s) (hts : t ⊆ s) : ContMDi protected theorem ContMDiff.contMDiffOn (hf : ContMDiff I I' n f) : ContMDiffOn I I' n f s := fun x _ => (hf x).contMDiffWithinAt -protected theorem Smooth.smoothOn (hf : Smooth I I' f) : SmoothOn I I' f s := - ContMDiff.contMDiffOn hf +@[deprecated (since := "2024-11-20")] alias Smooth.smoothOn := ContMDiff.contMDiffOn theorem contMDiffWithinAt_inter' (ht : t ∈ 𝓝[s] x) : ContMDiffWithinAt I I' n f (s ∩ t) x ↔ ContMDiffWithinAt I I' n f s x := @@ -716,16 +656,13 @@ protected theorem ContMDiffWithinAt.contMDiffAt ContMDiffAt I I' n f x := (contDiffWithinAt_localInvariantProp n).liftPropAt_of_liftPropWithinAt h ht -protected theorem SmoothWithinAt.smoothAt (h : SmoothWithinAt I I' f s x) (ht : s ∈ 𝓝 x) : - SmoothAt I I' f x := - ContMDiffWithinAt.contMDiffAt h ht +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.smoothAt := ContMDiffWithinAt.contMDiffAt protected theorem ContMDiffOn.contMDiffAt (h : ContMDiffOn I I' n f s) (hx : s ∈ 𝓝 x) : ContMDiffAt I I' n f x := (h x (mem_of_mem_nhds hx)).contMDiffAt hx -protected theorem SmoothOn.smoothAt (h : SmoothOn I I' f s) (hx : s ∈ 𝓝 x) : SmoothAt I I' f x := - h.contMDiffAt hx +@[deprecated (since := "2024-11-20")] alias SmoothOn.smoothAt := ContMDiffOn.contMDiffAt theorem contMDiffOn_iff_source_of_mem_maximalAtlas [SmoothManifoldWithCorners I M] (he : e ∈ maximalAtlas I M) (hs : s ⊆ e.source) : diff --git a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean index de822fbecd69c..e11bb6f43fe1c 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/NormedSpace.lean @@ -105,7 +105,8 @@ theorem ContinuousLinearMap.contMDiffWithinAt (L : E →L[𝕜] F) {s x} : theorem ContinuousLinearMap.contMDiffOn (L : E →L[𝕜] F) {s} : ContMDiffOn 𝓘(𝕜, E) 𝓘(𝕜, F) n L s := L.contMDiff.contMDiffOn -theorem ContinuousLinearMap.smooth (L : E →L[𝕜] F) : Smooth 𝓘(𝕜, E) 𝓘(𝕜, F) L := L.contMDiff +@[deprecated (since := "2024-11-20")] +alias ContinuousLinearMap.smooth := ContinuousLinearMap.contMDiff theorem ContMDiffWithinAt.clm_precomp {f : M → F₁ →L[𝕜] F₂} {s : Set M} {x : M} (hf : ContMDiffWithinAt I 𝓘(𝕜, F₁ →L[𝕜] F₂) n f s x) : @@ -261,13 +262,15 @@ theorem ContMDiff.clm_prodMap {g : M → F₁ →L[𝕜] F₃} {f : M → F₂ variable {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] /-- On any vector space, multiplication by a scalar is a smooth operation. -/ -theorem smooth_smul : Smooth (𝓘(𝕜).prod 𝓘(𝕜, V)) 𝓘(𝕜, V) fun p : 𝕜 × V => p.1 • p.2 := - smooth_iff.2 ⟨continuous_smul, fun _ _ => contDiff_smul.contDiffOn⟩ +theorem contMDiff_smul : ContMDiff (𝓘(𝕜).prod 𝓘(𝕜, V)) 𝓘(𝕜, V) ⊤ fun p : 𝕜 × V => p.1 • p.2 := + contMDiff_iff.2 ⟨continuous_smul, fun _ _ => contDiff_smul.contDiffOn⟩ + +@[deprecated (since := "2024-11-20")] alias smooth_smul := contMDiff_smul theorem ContMDiffWithinAt.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiffWithinAt I 𝓘(𝕜) n f s x) (hg : ContMDiffWithinAt I 𝓘(𝕜, V) n g s x) : ContMDiffWithinAt I 𝓘(𝕜, V) n (fun p => f p • g p) s x := - (smooth_smul.of_le le_top).contMDiffAt.comp_contMDiffWithinAt x (hf.prod_mk hg) + (contMDiff_smul.of_le le_top).contMDiffAt.comp_contMDiffWithinAt x (hf.prod_mk hg) nonrec theorem ContMDiffAt.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiffAt I 𝓘(𝕜) n f x) (hg : ContMDiffAt I 𝓘(𝕜, V) n g x) : ContMDiffAt I 𝓘(𝕜, V) n (fun p => f p • g p) x := @@ -281,18 +284,10 @@ theorem ContMDiff.smul {f : M → 𝕜} {g : M → V} (hf : ContMDiff I 𝓘( (hg : ContMDiff I 𝓘(𝕜, V) n g) : ContMDiff I 𝓘(𝕜, V) n fun p => f p • g p := fun x => (hf x).smul (hg x) -nonrec theorem SmoothWithinAt.smul {f : M → 𝕜} {g : M → V} (hf : SmoothWithinAt I 𝓘(𝕜) f s x) - (hg : SmoothWithinAt I 𝓘(𝕜, V) g s x) : SmoothWithinAt I 𝓘(𝕜, V) (fun p => f p • g p) s x := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.smul := ContMDiffWithinAt.smul -nonrec theorem SmoothAt.smul {f : M → 𝕜} {g : M → V} (hf : SmoothAt I 𝓘(𝕜) f x) - (hg : SmoothAt I 𝓘(𝕜, V) g x) : SmoothAt I 𝓘(𝕜, V) (fun p => f p • g p) x := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.smul := ContMDiffAt.smul -nonrec theorem SmoothOn.smul {f : M → 𝕜} {g : M → V} (hf : SmoothOn I 𝓘(𝕜) f s) - (hg : SmoothOn I 𝓘(𝕜, V) g s) : SmoothOn I 𝓘(𝕜, V) (fun p => f p • g p) s := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.smul := ContMDiffOn.smul -nonrec theorem Smooth.smul {f : M → 𝕜} {g : M → V} (hf : Smooth I 𝓘(𝕜) f) - (hg : Smooth I 𝓘(𝕜, V) g) : Smooth I 𝓘(𝕜, V) fun p => f p • g p := - hf.smul hg +@[deprecated (since := "2024-11-20")] alias Smooth.smul := ContMDiff.smul diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean index 7e220374d4678..5562db2312500 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Product.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Product.lean @@ -81,38 +81,22 @@ theorem ContMDiff.prod_mk_space {f : M → E'} {g : M → F'} (hf : ContMDiff I (hg : ContMDiff I 𝓘(𝕜, F') n g) : ContMDiff I 𝓘(𝕜, E' × F') n fun x => (f x, g x) := fun x => (hf x).prod_mk_space (hg x) -nonrec theorem SmoothWithinAt.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothWithinAt I I' f s x) - (hg : SmoothWithinAt I J' g s x) : SmoothWithinAt I (I'.prod J') (fun x => (f x, g x)) s x := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.prod_mk := ContMDiffWithinAt.prod_mk -nonrec theorem SmoothWithinAt.prod_mk_space {f : M → E'} {g : M → F'} - (hf : SmoothWithinAt I 𝓘(𝕜, E') f s x) (hg : SmoothWithinAt I 𝓘(𝕜, F') g s x) : - SmoothWithinAt I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) s x := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.prod_mk_space := ContMDiffWithinAt.prod_mk_space -nonrec theorem SmoothAt.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothAt I I' f x) - (hg : SmoothAt I J' g x) : SmoothAt I (I'.prod J') (fun x => (f x, g x)) x := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_mk := ContMDiffAt.prod_mk -nonrec theorem SmoothAt.prod_mk_space {f : M → E'} {g : M → F'} (hf : SmoothAt I 𝓘(𝕜, E') f x) - (hg : SmoothAt I 𝓘(𝕜, F') g x) : SmoothAt I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) x := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_mk_space := ContMDiffAt.prod_mk_space -nonrec theorem SmoothOn.prod_mk {f : M → M'} {g : M → N'} (hf : SmoothOn I I' f s) - (hg : SmoothOn I J' g s) : SmoothOn I (I'.prod J') (fun x => (f x, g x)) s := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_mk := ContMDiffOn.prod_mk -nonrec theorem SmoothOn.prod_mk_space {f : M → E'} {g : M → F'} (hf : SmoothOn I 𝓘(𝕜, E') f s) - (hg : SmoothOn I 𝓘(𝕜, F') g s) : SmoothOn I 𝓘(𝕜, E' × F') (fun x => (f x, g x)) s := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_mk_space := ContMDiffOn.prod_mk_space -nonrec theorem Smooth.prod_mk {f : M → M'} {g : M → N'} (hf : Smooth I I' f) (hg : Smooth I J' g) : - Smooth I (I'.prod J') fun x => (f x, g x) := - hf.prod_mk hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_mk := ContMDiff.prod_mk -nonrec theorem Smooth.prod_mk_space {f : M → E'} {g : M → F'} (hf : Smooth I 𝓘(𝕜, E') f) - (hg : Smooth I 𝓘(𝕜, F') g) : Smooth I 𝓘(𝕜, E' × F') fun x => (f x, g x) := - hf.prod_mk_space hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_mk_space := ContMDiff.prod_mk_space end ProdMk @@ -146,18 +130,13 @@ theorem contMDiffOn_fst {s : Set (M × N)} : ContMDiffOn (I.prod J) I n Prod.fst theorem contMDiff_fst : ContMDiff (I.prod J) I n (@Prod.fst M N) := fun _ => contMDiffAt_fst -theorem smoothWithinAt_fst {s : Set (M × N)} {p : M × N} : - SmoothWithinAt (I.prod J) I Prod.fst s p := - contMDiffWithinAt_fst +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_fst := contMDiffWithinAt_fst -theorem smoothAt_fst {p : M × N} : SmoothAt (I.prod J) I Prod.fst p := - contMDiffAt_fst +@[deprecated (since := "2024-11-20")] alias smoothAt_fst := contMDiffAt_fst -theorem smoothOn_fst {s : Set (M × N)} : SmoothOn (I.prod J) I Prod.fst s := - contMDiffOn_fst +@[deprecated (since := "2024-11-20")] alias smoothOn_fst := contMDiffOn_fst -theorem smooth_fst : Smooth (I.prod J) I (@Prod.fst M N) := - contMDiff_fst +@[deprecated (since := "2024-11-20")] alias smooth_fst := contMDiff_fst theorem ContMDiffAt.fst {f : N → M × M'} {x : N} (hf : ContMDiffAt J (I.prod I') n f x) : ContMDiffAt J I n (fun x => (f x).1) x := @@ -167,12 +146,9 @@ theorem ContMDiff.fst {f : N → M × M'} (hf : ContMDiff J (I.prod I') n f) : ContMDiff J I n fun x => (f x).1 := contMDiff_fst.comp hf -theorem SmoothAt.fst {f : N → M × M'} {x : N} (hf : SmoothAt J (I.prod I') f x) : - SmoothAt J I (fun x => (f x).1) x := - smoothAt_fst.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.fst := ContMDiffAt.fst -theorem Smooth.fst {f : N → M × M'} (hf : Smooth J (I.prod I') f) : Smooth J I fun x => (f x).1 := - smooth_fst.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.fst := ContMDiff.fst theorem contMDiffWithinAt_snd {s : Set (M × N)} {p : M × N} : ContMDiffWithinAt (I.prod J) J n Prod.snd s p := by @@ -202,18 +178,13 @@ theorem contMDiffOn_snd {s : Set (M × N)} : ContMDiffOn (I.prod J) J n Prod.snd theorem contMDiff_snd : ContMDiff (I.prod J) J n (@Prod.snd M N) := fun _ => contMDiffAt_snd -theorem smoothWithinAt_snd {s : Set (M × N)} {p : M × N} : - SmoothWithinAt (I.prod J) J Prod.snd s p := - contMDiffWithinAt_snd +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_snd := contMDiffWithinAt_snd -theorem smoothAt_snd {p : M × N} : SmoothAt (I.prod J) J Prod.snd p := - contMDiffAt_snd +@[deprecated (since := "2024-11-20")] alias smoothAt_snd := contMDiffAt_snd -theorem smoothOn_snd {s : Set (M × N)} : SmoothOn (I.prod J) J Prod.snd s := - contMDiffOn_snd +@[deprecated (since := "2024-11-20")] alias smoothOn_snd := contMDiffOn_snd -theorem smooth_snd : Smooth (I.prod J) J (@Prod.snd M N) := - contMDiff_snd +@[deprecated (since := "2024-11-20")] alias smooth_snd := contMDiff_snd theorem ContMDiffAt.snd {f : N → M × M'} {x : N} (hf : ContMDiffAt J (I.prod I') n f x) : ContMDiffAt J I' n (fun x => (f x).2) x := @@ -223,12 +194,9 @@ theorem ContMDiff.snd {f : N → M × M'} (hf : ContMDiff J (I.prod I') n f) : ContMDiff J I' n fun x => (f x).2 := contMDiff_snd.comp hf -theorem SmoothAt.snd {f : N → M × M'} {x : N} (hf : SmoothAt J (I.prod I') f x) : - SmoothAt J I' (fun x => (f x).2) x := - smoothAt_snd.comp x hf +@[deprecated (since := "2024-11-20")] alias SmoothAt.snd := ContMDiffAt.snd -theorem Smooth.snd {f : N → M × M'} (hf : Smooth J (I.prod I') f) : Smooth J I' fun x => (f x).2 := - smooth_snd.comp hf +@[deprecated (since := "2024-11-20")] alias Smooth.snd := ContMDiff.snd end Projections @@ -279,17 +247,16 @@ theorem contMDiff_prod_module_iff (f : M → F₁ × F₂) : rw [modelWithCornersSelf_prod, ← chartedSpaceSelf_prod] exact contMDiff_prod_iff f -theorem smoothAt_prod_iff (f : M → M' × N') {x : M} : - SmoothAt I (I'.prod J') f x ↔ SmoothAt I I' (Prod.fst ∘ f) x ∧ SmoothAt I J' (Prod.snd ∘ f) x := - contMDiffAt_prod_iff f +theorem contMDiff_prod_assoc : + ContMDiff ((I.prod I').prod J) (I.prod (I'.prod J)) n + fun x : (M × M') × N => (x.1.1, x.1.2, x.2) := + contMDiff_fst.fst.prod_mk <| contMDiff_fst.snd.prod_mk contMDiff_snd + +@[deprecated (since := "2024-11-20")] alias smoothAt_prod_iff := contMDiffAt_prod_iff -theorem smooth_prod_iff (f : M → M' × N') : - Smooth I (I'.prod J') f ↔ Smooth I I' (Prod.fst ∘ f) ∧ Smooth I J' (Prod.snd ∘ f) := - contMDiff_prod_iff f +@[deprecated (since := "2024-11-20")] alias smooth_prod_iff := contMDiff_prod_iff -theorem smooth_prod_assoc : - Smooth ((I.prod I').prod J) (I.prod (I'.prod J)) fun x : (M × M') × N => (x.1.1, x.1.2, x.2) := - smooth_fst.fst.prod_mk <| smooth_fst.snd.prod_mk smooth_snd +@[deprecated (since := "2024-11-20")] alias smooth_prod_assoc := contMDiff_prod_assoc section prodMap @@ -329,22 +296,13 @@ theorem ContMDiff.prod_map (hf : ContMDiff I I' n f) (hg : ContMDiff J J' n g) : intro p exact (hf p.1).prod_map' (hg p.2) -nonrec theorem SmoothWithinAt.prod_map (hf : SmoothWithinAt I I' f s x) - (hg : SmoothWithinAt J J' g r y) : - SmoothWithinAt (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) (x, y) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothWithinAt.prod_map := ContMDiffWithinAt.prod_map -nonrec theorem SmoothAt.prod_map (hf : SmoothAt I I' f x) (hg : SmoothAt J J' g y) : - SmoothAt (I.prod J) (I'.prod J') (Prod.map f g) (x, y) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothAt.prod_map := ContMDiffAt.prod_map -nonrec theorem SmoothOn.prod_map (hf : SmoothOn I I' f s) (hg : SmoothOn J J' g r) : - SmoothOn (I.prod J) (I'.prod J') (Prod.map f g) (s ×ˢ r) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias SmoothOn.prod_map := ContMDiffOn.prod_map -nonrec theorem Smooth.prod_map (hf : Smooth I I' f) (hg : Smooth J J' g) : - Smooth (I.prod J) (I'.prod J') (Prod.map f g) := - hf.prod_map hg +@[deprecated (since := "2024-11-20")] alias Smooth.prod_map := ContMDiff.prod_map end prodMap @@ -380,20 +338,12 @@ theorem contMDiff_pi_space : ContMDiff I 𝓘(𝕜, ∀ i, Fi i) n φ ↔ ∀ i, ContMDiff I 𝓘(𝕜, Fi i) n fun x => φ x i := ⟨fun h i x => contMDiffAt_pi_space.1 (h x) i, fun h x => contMDiffAt_pi_space.2 fun i => h i x⟩ -theorem smoothWithinAt_pi_space : - SmoothWithinAt I 𝓘(𝕜, ∀ i, Fi i) φ s x ↔ - ∀ i, SmoothWithinAt I 𝓘(𝕜, Fi i) (fun x => φ x i) s x := - contMDiffWithinAt_pi_space +@[deprecated (since := "2024-11-20")] alias smoothWithinAt_pi_space := contMDiffWithinAt_pi_space -theorem smoothOn_pi_space : - SmoothOn I 𝓘(𝕜, ∀ i, Fi i) φ s ↔ ∀ i, SmoothOn I 𝓘(𝕜, Fi i) (fun x => φ x i) s := - contMDiffOn_pi_space +@[deprecated (since := "2024-11-20")] alias smoothAt_pi_space := contMDiffAt_pi_space -theorem smoothAt_pi_space : - SmoothAt I 𝓘(𝕜, ∀ i, Fi i) φ x ↔ ∀ i, SmoothAt I 𝓘(𝕜, Fi i) (fun x => φ x i) x := - contMDiffAt_pi_space +@[deprecated (since := "2024-11-20")] alias smoothOn_pi_space := contMDiffOn_pi_space -theorem smooth_pi_space : Smooth I 𝓘(𝕜, ∀ i, Fi i) φ ↔ ∀ i, Smooth I 𝓘(𝕜, Fi i) fun x => φ x i := - contMDiff_pi_space +@[deprecated (since := "2024-11-20")] alias smooth_pi_space := contMDiff_pi_space end PiSpace diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index 9ab1a8b787d06..ad3c237062762 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -343,10 +343,7 @@ theorem ContMDiff.continuous_tangentMap (hf : ContMDiff I I' n f) (hmn : 1 ≤ n convert hf.continuousOn_tangentMapWithin hmn uniqueMDiffOn_univ rw [tangentMapWithin_univ] -/-- If a function is smooth, then its bundled derivative is smooth. -/ -theorem Smooth.tangentMap (hf : Smooth I I' f) : - Smooth I.tangent I'.tangent (tangentMap I I' f) := - ContMDiff.contMDiff_tangentMap hf le_rfl +@[deprecated (since := "2024-11-21")] alias Smooth.tangentMap := ContMDiff.contMDiff_tangentMap end tangentMap @@ -376,9 +373,9 @@ theorem tangentMap_tangentBundle_pure [Is : SmoothManifoldWithCorners I M] (p : · apply (PartialHomeomorph.open_target _).preimage I.continuous_invFun · simp only [mfld_simps] have A : MDifferentiableAt I I.tangent (fun x => @TotalSpace.mk M E (TangentSpace I) x 0) x := - haveI : Smooth I (I.prod 𝓘(𝕜, E)) (zeroSection E (TangentSpace I : M → Type _)) := - Bundle.smooth_zeroSection 𝕜 (TangentSpace I : M → Type _) - this.mdifferentiableAt + haveI : ContMDiff I (I.prod 𝓘(𝕜, E)) ⊤ (zeroSection E (TangentSpace I : M → Type _)) := + Bundle.contMDiff_zeroSection 𝕜 (TangentSpace I : M → Type _) + this.mdifferentiableAt le_top have B : fderivWithin 𝕜 (fun x' : E ↦ (x', (0 : E))) (Set.range I) (I ((chartAt H x) x)) v = (v, 0) := by rw [fderivWithin_eq_fderiv, DifferentiableAt.fderiv_prod] diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index 8566926007fa4..f8bd9f7980ad8 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -29,10 +29,7 @@ variable (I I') in def ContMDiffMap := { f : M → M' // ContMDiff I I' n f } -variable (I I') in -/-- Bundled smooth maps. -/ -abbrev SmoothMap := - ContMDiffMap I I' M M' ⊤ +@[deprecated (since := "024-11-21")] alias SmoothMap := ContMDiffMap @[inherit_doc] scoped[Manifold] notation "C^" n "⟮" I ", " M "; " I' ", " M' "⟯" => ContMDiffMap I I' M M' n @@ -54,8 +51,7 @@ instance instFunLike : FunLike C^n⟮I, M; I', M'⟯ M M' where protected theorem contMDiff (f : C^n⟮I, M; I', M'⟯) : ContMDiff I I' n f := f.prop -protected theorem smooth (f : C^∞⟮I, M; I', M'⟯) : Smooth I I' f := - f.prop +@[deprecated (since := "2024-11-20")] alias smooth := ContMDiffMap.contMDiff -- Porting note: use generic instance instead -- instance : Coe C^n⟮I, M; I', M'⟯ C(M, M') := diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 8a2f3e373a09a..5ee61a8605fe2 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -125,7 +125,7 @@ protected theorem contMDiffWithinAt (h : M ≃ₘ^n⟮I, I'⟯ M') {s x} : ContM protected theorem contDiff (h : E ≃ₘ^n⟮𝓘(𝕜, E), 𝓘(𝕜, E')⟯ E') : ContDiff 𝕜 n h := h.contMDiff.contDiff -protected theorem smooth (h : M ≃ₘ⟮I, I'⟯ M') : Smooth I I' h := h.contMDiff +@[deprecated (since := "2024-11-21")] alias smooth := Diffeomorph.contDiff protected theorem mdifferentiable (h : M ≃ₘ^n⟮I, I'⟯ M') (hn : 1 ≤ n) : MDifferentiable I I' h := h.contMDiff.mdifferentiable hn @@ -539,9 +539,8 @@ theorem contMDiff_transDiffeomorph_right {f : M' → M} : ContMDiff I' (I.transDiffeomorph e) n f ↔ ContMDiff I' I n f := (toTransDiffeomorph I M e).contMDiff_diffeomorph_comp_iff le_top -theorem smooth_transDiffeomorph_right {f : M' → M} : - Smooth I' (I.transDiffeomorph e) f ↔ Smooth I' I f := - contMDiff_transDiffeomorph_right e +@[deprecated (since := "2024-11-21")] +alias smooth_transDiffeomorph_right := contMDiff_transDiffeomorph_right @[simp] theorem contMDiffWithinAt_transDiffeomorph_left {f : M → M'} {x s} : @@ -563,8 +562,7 @@ theorem contMDiff_transDiffeomorph_left {f : M → M'} : ContMDiff (I.transDiffeomorph e) I' n f ↔ ContMDiff I I' n f := ((toTransDiffeomorph I M e).contMDiff_comp_diffeomorph_iff le_top).symm -theorem smooth_transDiffeomorph_left {f : M → M'} : - Smooth (I.transDiffeomorph e) I' f ↔ Smooth I I' f := - e.contMDiff_transDiffeomorph_left +@[deprecated (since := "2024-11-21")] +alias smooth_transDiffeomorph_left := contMDiff_transDiffeomorph_left end Diffeomorph diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 552b4986092f8..377ab9d17b795 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -477,27 +477,30 @@ theorem ContMDiff.mdifferentiableAt (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : MDifferentiableAt I I' f x := hf.contMDiffAt.mdifferentiableAt hn +theorem ContMDiff.mdifferentiableWithinAt (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : + MDifferentiableWithinAt I I' f s x := + (hf.contMDiffAt.mdifferentiableAt hn).mdifferentiableWithinAt + theorem ContMDiffOn.mdifferentiableOn (hf : ContMDiffOn I I' n f s) (hn : 1 ≤ n) : MDifferentiableOn I I' f s := fun x hx => (hf x hx).mdifferentiableWithinAt hn -nonrec theorem SmoothWithinAt.mdifferentiableWithinAt (hf : SmoothWithinAt I I' f s x) : - MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableWithinAt le_top +@[deprecated (since := "2024-11-20")] +alias SmoothWithinAt.mdifferentiableWithinAt := ContMDiffWithinAt.mdifferentiableWithinAt theorem ContMDiff.mdifferentiable (hf : ContMDiff I I' n f) (hn : 1 ≤ n) : MDifferentiable I I' f := fun x => (hf x).mdifferentiableAt hn -nonrec theorem SmoothAt.mdifferentiableAt (hf : SmoothAt I I' f x) : MDifferentiableAt I I' f x := - hf.mdifferentiableAt le_top +@[deprecated (since := "2024-11-20")] +alias SmoothAt.mdifferentiableAt := ContMDiffAt.mdifferentiableAt -nonrec theorem SmoothOn.mdifferentiableOn (hf : SmoothOn I I' f s) : MDifferentiableOn I I' f s := - hf.mdifferentiableOn le_top +@[deprecated (since := "2024-11-20")] +alias SmoothOn.mdifferentiableOn := ContMDiffOn.mdifferentiableOn -theorem Smooth.mdifferentiable (hf : Smooth I I' f) : MDifferentiable I I' f := - ContMDiff.mdifferentiable hf le_top +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiable := ContMDiff.mdifferentiable -theorem Smooth.mdifferentiableAt (hf : Smooth I I' f) : MDifferentiableAt I I' f x := - hf.mdifferentiable x +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiableAt := ContMDiff.mdifferentiableAt theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : ContinuousOn f s := fun x hx => (h x hx).continuousWithinAt @@ -505,8 +508,8 @@ theorem MDifferentiableOn.continuousOn (h : MDifferentiableOn I I' f s) : Contin theorem MDifferentiable.continuous (h : MDifferentiable I I' f) : Continuous f := continuous_iff_continuousAt.2 fun x => (h x).continuousAt -theorem Smooth.mdifferentiableWithinAt (hf : Smooth I I' f) : MDifferentiableWithinAt I I' f s x := - hf.mdifferentiableAt.mdifferentiableWithinAt +@[deprecated (since := "2024-11-20")] +alias Smooth.mdifferentiableWithinAt := ContMDiff.mdifferentiableWithinAt /-! ### Deriving continuity from differentiability on manifolds -/ diff --git a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean index 4facc7f820910..e090b101ab9e7 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/UniqueDifferential.lean @@ -120,7 +120,7 @@ variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {Z : M → Type theorem Trivialization.mdifferentiable (e : Trivialization F (π F Z)) [MemTrivializationAtlas e] : e.toPartialHomeomorph.MDifferentiable (I.prod 𝓘(𝕜, F)) (I.prod 𝓘(𝕜, F)) := - ⟨e.smoothOn.mdifferentiableOn, e.smoothOn_symm.mdifferentiableOn⟩ + ⟨e.contMDiffOn.mdifferentiableOn le_top, e.contMDiffOn_symm.mdifferentiableOn le_top⟩ theorem UniqueMDiffWithinAt.smooth_bundle_preimage {p : TotalSpace F Z} (hs : UniqueMDiffWithinAt I s p.proj) : diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 35ffa8cfc8743..3d001e4e01c5d 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -160,8 +160,10 @@ theorem sum_le_one (x : M) : ∑ᶠ i, f i x ≤ 1 := def toPartitionOfUnity : PartitionOfUnity ι M s := { f with toFun := fun i => f i } -theorem smooth_sum : Smooth I 𝓘(ℝ) fun x => ∑ᶠ i, f i x := - smooth_finsum (fun i => (f i).smooth) f.locallyFinite +theorem contMDiff_sum : ContMDiff I 𝓘(ℝ) ⊤ fun x => ∑ᶠ i, f i x := + contMDiff_finsum (fun i => (f i).contMDiff) f.locallyFinite + +@[deprecated (since := "2024-11-21")] alias smooth_sum := contMDiff_sum theorem le_one (i : ι) (x : M) : f i x ≤ 1 := f.toPartitionOfUnity.le_one i x @@ -178,9 +180,7 @@ theorem contMDiff_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), ContMDi contMDiff_of_tsupport fun x hx => ((f i).contMDiff.contMDiffAt.of_le le_top).smul <| hg x <| tsupport_smul_subset_left _ _ hx -theorem smooth_smul {g : M → F} {i} (hg : ∀ x ∈ tsupport (f i), SmoothAt I 𝓘(ℝ, F) g x) : - Smooth I 𝓘(ℝ, F) fun x => f i x • g x := - f.contMDiff_smul hg +@[deprecated (since := "2024-11-21")] alias smooth_smul := contMDiff_smul /-- If `f` is a smooth partition of unity on a set `s : Set M` and `g : ι → M → F` is a family of functions such that `g i` is $C^n$ smooth at every point of the topological support of `f i`, then @@ -191,20 +191,14 @@ theorem contMDiff_finsum_smul {g : ι → M → F} (contMDiff_finsum fun i => f.contMDiff_smul (hg i)) <| f.locallyFinite.subset fun _ => support_smul_subset_left _ _ -/-- If `f` is a smooth partition of unity on a set `s : Set M` and `g : ι → M → F` is a family of -functions such that `g i` is smooth at every point of the topological support of `f i`, then the sum -`fun x ↦ ∑ᶠ i, f i x • g i x` is smooth on the whole manifold. -/ -theorem smooth_finsum_smul {g : ι → M → F} - (hg : ∀ (i), ∀ x ∈ tsupport (f i), SmoothAt I 𝓘(ℝ, F) (g i) x) : - Smooth I 𝓘(ℝ, F) fun x => ∑ᶠ i, f i x • g i x := - f.contMDiff_finsum_smul hg +@[deprecated (since := "2024-11-21")] alias smooth_finsum_smul := contMDiff_finsum_smul theorem contMDiffAt_finsum {x₀ : M} {g : ι → M → F} (hφ : ∀ i, x₀ ∈ tsupport (f i) → ContMDiffAt I 𝓘(ℝ, F) n (g i) x₀) : ContMDiffAt I 𝓘(ℝ, F) n (fun x ↦ ∑ᶠ i, f i x • g i x) x₀ := by refine _root_.contMDiffAt_finsum (f.locallyFinite.smul_left _) fun i ↦ ?_ by_cases hx : x₀ ∈ tsupport (f i) - · exact ContMDiffAt.smul ((f i).smooth.of_le le_top).contMDiffAt (hφ i hx) + · exact ContMDiffAt.smul ((f i).contMDiff.of_le le_top).contMDiffAt (hφ i hx) · exact contMDiffAt_of_not_mem (compl_subset_compl.mpr (tsupport_smul_subset_left (f i) (g i)) hx) n @@ -294,13 +288,8 @@ theorem IsSubordinate.contMDiff_finsum_smul {g : ι → M → F} (hf : f.IsSubor ContMDiff I 𝓘(ℝ, F) n fun x => ∑ᶠ i, f i x • g i x := f.contMDiff_finsum_smul fun i _ hx => (hg i).contMDiffAt <| (ho i).mem_nhds (hf i hx) -/-- If `f` is a smooth partition of unity on a set `s : Set M` subordinate to a family of open sets -`U : ι → Set M` and `g : ι → M → F` is a family of functions such that `g i` is smooth on `U i`, -then the sum `fun x ↦ ∑ᶠ i, f i x • g i x` is smooth on the whole manifold. -/ -theorem IsSubordinate.smooth_finsum_smul {g : ι → M → F} (hf : f.IsSubordinate U) - (ho : ∀ i, IsOpen (U i)) (hg : ∀ i, SmoothOn I 𝓘(ℝ, F) (g i) (U i)) : - Smooth I 𝓘(ℝ, F) fun x => ∑ᶠ i, f i x • g i x := - hf.contMDiff_finsum_smul ho hg +@[deprecated (since := "2024-11-21")] +alias IsSubordinate.smooth_finsum_smul := IsSubordinate.contMDiff_finsum_smul end IsSubordinate @@ -309,14 +298,17 @@ end SmoothPartitionOfUnity namespace BumpCovering -- Repeat variables to drop `[FiniteDimensional ℝ E]` and `[SmoothManifoldWithCorners I M]` -theorem smooth_toPartitionOfUnity {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] +theorem contMDiff_toPartitionOfUnity {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E] {H : Type uH} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} {M : Type uM} [TopologicalSpace M] [ChartedSpace H M] {s : Set M} (f : BumpCovering ι M s) - (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) (i : ι) : Smooth I 𝓘(ℝ) (f.toPartitionOfUnity i) := - (hf i).mul <| (smooth_finprod_cond fun j _ => smooth_const.sub (hf j)) <| by + (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) (i : ι) : ContMDiff I 𝓘(ℝ) ⊤ (f.toPartitionOfUnity i) := + (hf i).mul <| (contMDiff_finprod_cond fun j _ => contMDiff_const.sub (hf j)) <| by simp only [Pi.sub_def, mulSupport_one_sub] exact f.locallyFinite +@[deprecated (since := "2024-11-21")] +alias smooth_toPartitionOfUnity := contMDiff_toPartitionOfUnity + variable {s : Set M} /-- A `BumpCovering` such that all functions in this covering are smooth generates a smooth @@ -326,24 +318,24 @@ In our formalization, not every `f : BumpCovering ι M s` with smooth functions `SmoothBumpCovering`; instead, a `SmoothBumpCovering` is a covering by supports of `SmoothBumpFunction`s. So, we define `BumpCovering.toSmoothPartitionOfUnity`, then reuse it in `SmoothBumpCovering.toSmoothPartitionOfUnity`. -/ -def toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : +def toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : SmoothPartitionOfUnity ι I M s := { f.toPartitionOfUnity with - toFun := fun i => ⟨f.toPartitionOfUnity i, f.smooth_toPartitionOfUnity hf i⟩ } + toFun := fun i => ⟨f.toPartitionOfUnity i, f.contMDiff_toPartitionOfUnity hf i⟩ } @[simp] theorem toSmoothPartitionOfUnity_toPartitionOfUnity (f : BumpCovering ι M s) - (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : + (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : (f.toSmoothPartitionOfUnity hf).toPartitionOfUnity = f.toPartitionOfUnity := rfl @[simp] -theorem coe_toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) +theorem coe_toSmoothPartitionOfUnity (f : BumpCovering ι M s) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) (i : ι) : ⇑(f.toSmoothPartitionOfUnity hf i) = f.toPartitionOfUnity i := rfl theorem IsSubordinate.toSmoothPartitionOfUnity {f : BumpCovering ι M s} {U : ι → Set M} - (h : f.IsSubordinate U) (hf : ∀ i, Smooth I 𝓘(ℝ) (f i)) : + (h : f.IsSubordinate U) (hf : ∀ i, ContMDiff I 𝓘(ℝ) ⊤ (f i)) : (f.toSmoothPartitionOfUnity hf).IsSubordinate U := h.toPartitionOfUnity @@ -457,7 +449,7 @@ alias ⟨_, IsSubordinate.toBumpCovering⟩ := isSubordinate_toBumpCovering /-- Every `SmoothBumpCovering` defines a smooth partition of unity. -/ def toSmoothPartitionOfUnity : SmoothPartitionOfUnity ι I M s := - fs.toBumpCovering.toSmoothPartitionOfUnity fun i => (fs i).smooth + fs.toBumpCovering.toSmoothPartitionOfUnity fun i => (fs i).contMDiff theorem toSmoothPartitionOfUnity_apply (i : ι) (x : M) : fs.toSmoothPartitionOfUnity i x = fs i x * ∏ᶠ (j) (_ : WellOrderingRel j i), (1 - fs j x) := @@ -511,7 +503,7 @@ theorem exists_smooth_zero_one_of_isClosed [T2Space M] [SigmaCompactSpace M] {s rcases SmoothBumpCovering.exists_isSubordinate I ht this with ⟨ι, f, hf⟩ set g := f.toSmoothPartitionOfUnity refine - ⟨⟨_, g.smooth_sum⟩, fun x hx => ?_, fun x => g.sum_eq_one, fun x => + ⟨⟨_, g.contMDiff_sum⟩, fun x hx => ?_, fun x => g.sum_eq_one, fun x => ⟨g.sum_nonneg x, g.sum_le_one x⟩⟩ suffices ∀ i, g i x = 0 by simp only [this, ContMDiffMap.coeFn_mk, finsum_zero, Pi.zero_apply] refine fun i => f.toSmoothPartitionOfUnity_zero_of_zero ?_ @@ -554,8 +546,9 @@ def single (i : ι) (s : Set M) : SmoothPartitionOfUnity ι I M s := (BumpCovering.single i s).toSmoothPartitionOfUnity fun j => by classical rcases eq_or_ne j i with (rfl | h) - · simp only [smooth_one, ContinuousMap.coe_one, BumpCovering.coe_single, Pi.single_eq_same] - · simp only [smooth_zero, BumpCovering.coe_single, Pi.single_eq_of_ne h, ContinuousMap.coe_zero] + · simp only [contMDiff_one, ContinuousMap.coe_one, BumpCovering.coe_single, Pi.single_eq_same] + · simp only [contMDiff_zero, BumpCovering.coe_single, Pi.single_eq_of_ne h, + ContinuousMap.coe_zero] instance [Inhabited ι] (s : Set M) : Inhabited (SmoothPartitionOfUnity ι I M s) := ⟨single I default s⟩ @@ -570,12 +563,12 @@ theorem exists_isSubordinate {s : Set M} (hs : IsClosed s) (U : ι → Set M) (h haveI : LocallyCompactSpace M := ChartedSpace.locallyCompactSpace H M -- porting note(https://github.com/leanprover/std4/issues/116): -- split `rcases` into `have` + `rcases` - have := BumpCovering.exists_isSubordinate_of_prop (Smooth I 𝓘(ℝ)) ?_ hs U ho hU + have := BumpCovering.exists_isSubordinate_of_prop (ContMDiff I 𝓘(ℝ) ⊤) ?_ hs U ho hU · rcases this with ⟨f, hf, hfU⟩ exact ⟨f.toSmoothPartitionOfUnity hf, hfU.toSmoothPartitionOfUnity hf⟩ · intro s t hs ht hd rcases exists_smooth_zero_one_of_isClosed I hs ht hd with ⟨f, hf⟩ - exact ⟨f, f.smooth, hf⟩ + exact ⟨f, f.contMDiff, hf⟩ theorem exists_isSubordinate_chartAt_source_of_isClosed {s : Set M} (hs : IsClosed s) : ∃ f : SmoothPartitionOfUnity s I M s, @@ -618,7 +611,7 @@ Then there exists a smooth function `g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯` such See also `exists_contMDiffOn_forall_mem_convex_of_local` and `exists_smooth_forall_mem_convex_of_local_const`. -/ theorem exists_smooth_forall_mem_convex_of_local (ht : ∀ x, Convex ℝ (t x)) - (Hloc : ∀ x : M, ∃ U ∈ 𝓝 x, ∃ g : M → F, SmoothOn I 𝓘(ℝ, F) g U ∧ ∀ y ∈ U, g y ∈ t y) : + (Hloc : ∀ x : M, ∃ U ∈ 𝓝 x, ∃ g : M → F, ContMDiffOn I 𝓘(ℝ, F) ⊤ g U ∧ ∀ y ∈ U, g y ∈ t y) : ∃ g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯, ∀ x, g x ∈ t x := exists_contMDiffOn_forall_mem_convex_of_local I ht Hloc @@ -631,7 +624,7 @@ theorem exists_smooth_forall_mem_convex_of_local_const (ht : ∀ x, Convex ℝ ( (Hloc : ∀ x : M, ∃ c : F, ∀ᶠ y in 𝓝 x, c ∈ t y) : ∃ g : C^∞⟮I, M; 𝓘(ℝ, F), F⟯, ∀ x, g x ∈ t x := exists_smooth_forall_mem_convex_of_local I ht fun x => let ⟨c, hc⟩ := Hloc x - ⟨_, hc, fun _ => c, smoothOn_const, fun _ => id⟩ + ⟨_, hc, fun _ => c, contMDiffOn_const, fun _ => id⟩ /-- Let `M` be a smooth σ-compact manifold with extended distance. Let `K : ι → Set M` be a locally finite family of closed sets, let `U : ι → Set M` be a family of open sets such that `K i ⊆ U i` for @@ -664,7 +657,7 @@ theorem Metric.exists_smooth_forall_closedBall_subset {M} [MetricSpace M] [Chart exact hδ i x hx lemma IsOpen.exists_msmooth_support_eq_aux {s : Set H} (hs : IsOpen s) : - ∃ f : H → ℝ, f.support = s ∧ Smooth I 𝓘(ℝ) f ∧ Set.range f ⊆ Set.Icc 0 1 := by + ∃ f : H → ℝ, f.support = s ∧ ContMDiff I 𝓘(ℝ) ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := by have h's : IsOpen (I.symm ⁻¹' s) := I.continuous_symm.isOpen_preimage _ hs rcases h's.exists_smooth_support_eq with ⟨f, f_supp, f_diff, f_range⟩ refine ⟨f ∘ I, ?_, ?_, ?_⟩ @@ -676,11 +669,11 @@ lemma IsOpen.exists_msmooth_support_eq_aux {s : Set H} (hs : IsOpen s) : /-- Given an open set in a finite-dimensional real manifold, there exists a nonnegative smooth function with support equal to `s`. -/ theorem IsOpen.exists_msmooth_support_eq {s : Set M} (hs : IsOpen s) : - ∃ f : M → ℝ, f.support = s ∧ Smooth I 𝓘(ℝ) f ∧ ∀ x, 0 ≤ f x := by + ∃ f : M → ℝ, f.support = s ∧ ContMDiff I 𝓘(ℝ) ⊤ f ∧ ∀ x, 0 ≤ f x := by rcases SmoothPartitionOfUnity.exists_isSubordinate_chartAt_source I M with ⟨f, hf⟩ have A : ∀ (c : M), ∃ g : H → ℝ, g.support = (chartAt H c).target ∩ (chartAt H c).symm ⁻¹' s ∧ - Smooth I 𝓘(ℝ) g ∧ Set.range g ⊆ Set.Icc 0 1 := by + ContMDiff I 𝓘(ℝ) ⊤ g ∧ Set.range g ⊆ Set.Icc 0 1 := by intro i apply IsOpen.exists_msmooth_support_eq_aux exact PartialHomeomorph.isOpen_inter_preimage_symm _ hs @@ -714,7 +707,7 @@ theorem IsOpen.exists_msmooth_support_eq {s : Set M} (hs : IsOpen s) : · have : x ∉ support (f c) := by contrapose! Hx; exact subset_tsupport _ Hx rw [nmem_support] at this simp [this] - · apply SmoothPartitionOfUnity.smooth_finsum_smul + · apply SmoothPartitionOfUnity.contMDiff_finsum_smul intro c x hx apply (g_diff c (chartAt H c x)).comp exact contMDiffAt_of_mem_maximalAtlas (SmoothManifoldWithCorners.chart_mem_maximalAtlas _) @@ -727,7 +720,7 @@ exists a smooth function with support equal to `s`, taking values in `[0,1]`, an exactly on `t`. -/ theorem exists_msmooth_support_eq_eq_one_iff {s t : Set M} (hs : IsOpen s) (ht : IsClosed t) (h : t ⊆ s) : - ∃ f : M → ℝ, Smooth I 𝓘(ℝ) f ∧ range f ⊆ Icc 0 1 ∧ support f = s + ∃ f : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ f ∧ range f ⊆ Icc 0 1 ∧ support f = s ∧ (∀ x, x ∈ t ↔ f x = 1) := by /- Take `f` with support equal to `s`, and `g` with support equal to `tᶜ`. Then `f / (f + g)` satisfies the conclusion of the theorem. -/ @@ -765,7 +758,7 @@ there exists an infinitely smooth function that is equal to `0` exactly on `s` a exactly on `t`. See also `exists_smooth_zero_one_of_isClosed` for a slightly weaker version. -/ theorem exists_msmooth_zero_iff_one_iff_of_isClosed {s t : Set M} (hs : IsClosed s) (ht : IsClosed t) (hd : Disjoint s t) : - ∃ f : M → ℝ, Smooth I 𝓘(ℝ) f ∧ range f ⊆ Icc 0 1 ∧ (∀ x, x ∈ s ↔ f x = 0) + ∃ f : M → ℝ, ContMDiff I 𝓘(ℝ) ⊤ f ∧ range f ⊆ Icc 0 1 ∧ (∀ x, x ∈ s ↔ f x = 0) ∧ (∀ x, x ∈ t ↔ f x = 1) := by rcases exists_msmooth_support_eq_eq_one_iff I hs.isOpen_compl ht hd.subset_compl_left with ⟨f, f_diff, f_range, fs, ft⟩ diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index 889506e8609ea..48aec60a959fa 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -97,7 +97,7 @@ theorem smoothSheafCommRing.isUnit_stalk_iff {x : M} #adaptation_note /-- https://github.com/leanprover/lean4/pull/6024 was `exact`; somehow `convert` bypasess unification issues -/ convert ((contDiffAt_inv _ (hVf y)).contMDiffAt).comp y - (f.smooth.comp (smooth_inclusion hUV)).smoothAt + (f.contMDiff.comp (contMDiff_inclusion hUV)).contMDiffAt /-- The non-units of the stalk at `x` of the sheaf of smooth functions from `M` to `𝕜`, considered as a sheaf of commutative rings, are the functions whose values at `x` are zero. -/ diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index b25eee8e25950..0f528359115a1 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -126,7 +126,7 @@ def smoothSheaf.evalAt (x : TopCat.of M) (U : OpenNhds x) lemma smoothSheaf.eval_surjective (x : M) : Function.Surjective (smoothSheaf.eval IM I N x) := by apply TopCat.stalkToFiber_surjective intro n - exact ⟨⊤, fun _ ↦ n, smooth_const, rfl⟩ + exact ⟨⊤, fun _ ↦ n, contMDiff_const, rfl⟩ instance [Nontrivial N] (x : M) : Nontrivial ((smoothSheaf IM I M N).presheaf.stalk x) := (smoothSheaf.eval_surjective IM I N x).nontrivial @@ -138,11 +138,14 @@ variable {IM I N} smoothSheaf.eval IM I N (x : M) ((smoothSheaf IM I M N).presheaf.germ U x hx f) = f ⟨x, hx⟩ := TopCat.stalkToFiber_germ ((contDiffWithinAt_localInvariantProp ⊤).localPredicate M N) _ _ _ _ -lemma smoothSheaf.smooth_section {U : (Opens (TopCat.of M))ᵒᵖ} +lemma smoothSheaf.contMDiff_section {U : (Opens (TopCat.of M))ᵒᵖ} (f : (smoothSheaf IM I M N).presheaf.obj U) : - Smooth IM I f := + ContMDiff IM I ⊤ f := (contDiffWithinAt_localInvariantProp ⊤).section_spec _ _ _ _ +@[deprecated (since := "2024-11-21")] +alias smoothSheaf.smooth_section := smoothSheaf.contMDiff_section + end TypeCat section LieGroup @@ -213,7 +216,7 @@ noncomputable def smoothSheafCommGroup : TopCat.Sheaf CommGrp.{u} (TopCat.of M) @[to_additive "For a manifold `M` and a smooth homomorphism `φ` between abelian additive Lie groups `A`, `A'`, the 'left-composition-by-`φ`' morphism of sheaves from `smoothSheafAddCommGroup IM I M A` to `smoothSheafAddCommGroup IM I' M A'`."] -def smoothSheafCommGroup.compLeft (φ : A →* A') (hφ : Smooth I I' φ) : +def smoothSheafCommGroup.compLeft (φ : A →* A') (hφ : ContMDiff I I' ⊤ φ) : smoothSheafCommGroup IM I M A ⟶ smoothSheafCommGroup IM I' M A' := CategoryTheory.Sheaf.Hom.mk <| { app := fun _ ↦ CommGrp.ofHom <| SmoothMap.compLeftMonoidHom _ _ φ hφ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean index dab21687a47a9..3fd4222c1daba 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Basic.lean @@ -212,34 +212,29 @@ theorem contMDiff_proj : ContMDiff (IB.prod 𝓘(𝕜, F)) IB n (π F E) := fun rw [contMDiffAt_totalSpace] at this exact this.1 -theorem smooth_proj : Smooth (IB.prod 𝓘(𝕜, F)) IB (π F E) := - contMDiff_proj E +@[deprecated (since := "2024-11-21")] alias smooth_proj := contMDiff_proj theorem contMDiffOn_proj {s : Set (TotalSpace F E)} : ContMDiffOn (IB.prod 𝓘(𝕜, F)) IB n (π F E) s := (Bundle.contMDiff_proj E).contMDiffOn -theorem smoothOn_proj {s : Set (TotalSpace F E)} : SmoothOn (IB.prod 𝓘(𝕜, F)) IB (π F E) s := - contMDiffOn_proj E +@[deprecated (since := "2024-11-21")] alias smoothOn_proj := contMDiffOn_proj theorem contMDiffAt_proj {p : TotalSpace F E} : ContMDiffAt (IB.prod 𝓘(𝕜, F)) IB n (π F E) p := (Bundle.contMDiff_proj E).contMDiffAt -theorem smoothAt_proj {p : TotalSpace F E} : SmoothAt (IB.prod 𝓘(𝕜, F)) IB (π F E) p := - Bundle.contMDiffAt_proj E +@[deprecated (since := "2024-11-21")] alias smoothAt_proj := contMDiffAt_proj theorem contMDiffWithinAt_proj {s : Set (TotalSpace F E)} {p : TotalSpace F E} : ContMDiffWithinAt (IB.prod 𝓘(𝕜, F)) IB n (π F E) s p := (Bundle.contMDiffAt_proj E).contMDiffWithinAt -theorem smoothWithinAt_proj {s : Set (TotalSpace F E)} {p : TotalSpace F E} : - SmoothWithinAt (IB.prod 𝓘(𝕜, F)) IB (π F E) s p := - Bundle.contMDiffWithinAt_proj E +@[deprecated (since := "2024-11-21")] alias smoothWithinAt_proj := contMDiffWithinAt_proj variable (𝕜) [∀ x, AddCommMonoid (E x)] variable [∀ x, Module 𝕜 (E x)] [VectorBundle 𝕜 F E] -theorem smooth_zeroSection : Smooth IB (IB.prod 𝓘(𝕜, F)) (zeroSection F E) := fun x ↦ by +theorem contMDiff_zeroSection : ContMDiff IB (IB.prod 𝓘(𝕜, F)) ⊤ (zeroSection F E) := fun x ↦ by unfold zeroSection rw [Bundle.contMDiffAt_section] apply (contMDiffAt_const (c := 0)).congr_of_eventuallyEq @@ -247,6 +242,8 @@ theorem smooth_zeroSection : Smooth IB (IB.prod 𝓘(𝕜, F)) (zeroSection F E) (mem_baseSet_trivializationAt F E x)] with y hy using congr_arg Prod.snd <| (trivializationAt F E x).zeroSection 𝕜 hy +@[deprecated (since := "2024-11-21")] alias smooth_zeroSection := contMDiff_zeroSection + end Bundle end @@ -272,9 +269,9 @@ topological vector bundle over `B` with fibers isomorphic to `F`, then `SmoothVe registers that the bundle is smooth, in the sense of having smooth transition functions. This is a mixin, not carrying any new data. -/ class SmoothVectorBundle : Prop where - protected smoothOn_coordChangeL : + protected contMDiffOn_coordChangeL : ∀ (e e' : Trivialization F (π F E)) [MemTrivializationAtlas e] [MemTrivializationAtlas e'], - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) variable [SmoothVectorBundle F E IB] @@ -284,27 +281,23 @@ section SmoothCoordChange variable {F E} variable (e e' : Trivialization F (π F E)) [MemTrivializationAtlas e] [MemTrivializationAtlas e'] -theorem smoothOn_coordChangeL : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) +theorem contMDiffOn_coordChangeL : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := - SmoothVectorBundle.smoothOn_coordChangeL e e' + (SmoothVectorBundle.contMDiffOn_coordChangeL e e').of_le le_top -theorem smoothOn_symm_coordChangeL : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) +theorem contMDiffOn_symm_coordChangeL : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) (e.baseSet ∩ e'.baseSet) := by + apply ContMDiffOn.of_le _ le_top rw [inter_comm] - refine (SmoothVectorBundle.smoothOn_coordChangeL e' e).congr fun b hb ↦ ?_ + refine (SmoothVectorBundle.contMDiffOn_coordChangeL e' e).congr fun b hb ↦ ?_ rw [e.symm_coordChangeL e' hb] -theorem contMDiffOn_coordChangeL : - ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) - (e.baseSet ∩ e'.baseSet) := - (smoothOn_coordChangeL e e').of_le le_top +@[deprecated (since := "2024-11-21")] alias smoothOn_coordChangeL := contMDiffOn_coordChangeL +@[deprecated (since := "2024-11-21")] +alias smoothOn_symm_coordChangeL := contMDiffOn_symm_coordChangeL -theorem contMDiffOn_symm_coordChangeL : - ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) n (fun b : B => ((e.coordChangeL 𝕜 e' b).symm : F →L[𝕜] F)) - (e.baseSet ∩ e'.baseSet) := - (smoothOn_symm_coordChangeL e e').of_le le_top variable {e e'} @@ -313,9 +306,7 @@ theorem contMDiffAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.ba (contMDiffOn_coordChangeL e e').contMDiffAt <| (e.open_baseSet.inter e'.open_baseSet).mem_nhds ⟨h, h'⟩ -theorem smoothAt_coordChangeL {x : B} (h : x ∈ e.baseSet) (h' : x ∈ e'.baseSet) : - SmoothAt IB 𝓘(𝕜, F →L[𝕜] F) (fun b : B => (e.coordChangeL 𝕜 e' b : F →L[𝕜] F)) x := - contMDiffAt_coordChangeL h h' +@[deprecated (since := "2024-11-21")] alias smoothAt_coordChangeL := contMDiffAt_coordChangeL variable {s : Set M} {f : M → B} {g : M → F} {x : M} @@ -339,25 +330,17 @@ protected theorem ContMDiff.coordChangeL ContMDiff IM 𝓘(𝕜, F →L[𝕜] F) n (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) := fun x ↦ (hf x).coordChangeL (he x) (he' x) -protected nonrec theorem SmoothWithinAt.coordChangeL - (hf : SmoothWithinAt IM IB f s x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothWithinAt IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) s x := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothWithinAt.coordChangeL := ContMDiffWithinAt.coordChangeL -protected nonrec theorem SmoothAt.coordChangeL - (hf : SmoothAt IM IB f x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothAt IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) x := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothAt.coordChangeL := ContMDiffAt.coordChangeL -protected nonrec theorem SmoothOn.coordChangeL - (hf : SmoothOn IM IB f s) (he : MapsTo f s e.baseSet) (he' : MapsTo f s e'.baseSet) : - SmoothOn IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) s := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias SmoothOn.coordChangeL := ContMDiffOn.coordChangeL -protected nonrec theorem Smooth.coordChangeL - (hf : Smooth IM IB f) (he : ∀ x, f x ∈ e.baseSet) (he' : ∀ x, f x ∈ e'.baseSet) : - Smooth IM 𝓘(𝕜, F →L[𝕜] F) (fun y ↦ (e.coordChangeL 𝕜 e' (f y) : F →L[𝕜] F)) := - hf.coordChangeL he he' +@[deprecated (since := "2024-11-21")] +alias Smooth.coordChangeL := ContMDiff.coordChangeL protected theorem ContMDiffWithinAt.coordChange (hf : ContMDiffWithinAt IM IB n f s x) (hg : ContMDiffWithinAt IM 𝓘(𝕜, F) n g s x) @@ -386,26 +369,17 @@ protected theorem ContMDiff.coordChange (hf : ContMDiff IM IB n f) ContMDiff IM 𝓘(𝕜, F) n (fun y ↦ e.coordChange e' (f y) (g y)) := fun x ↦ (hf x).coordChange (hg x) (he x) (he' x) -protected nonrec theorem SmoothWithinAt.coordChange - (hf : SmoothWithinAt IM IB f s x) (hg : SmoothWithinAt IM 𝓘(𝕜, F) g s x) - (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothWithinAt IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) s x := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothWithinAt.coordChange := ContMDiffWithinAt.coordChange -protected nonrec theorem SmoothAt.coordChange (hf : SmoothAt IM IB f x) - (hg : SmoothAt IM 𝓘(𝕜, F) g x) (he : f x ∈ e.baseSet) (he' : f x ∈ e'.baseSet) : - SmoothAt IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) x := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothAt.coordChange := ContMDiffAt.coordChange -protected nonrec theorem SmoothOn.coordChange (hf : SmoothOn IM IB f s) - (hg : SmoothOn IM 𝓘(𝕜, F) g s) (he : MapsTo f s e.baseSet) (he' : MapsTo f s e'.baseSet) : - SmoothOn IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) s := - hf.coordChange hg he he' +@[deprecated (since := "2024-11-21")] +alias SmoothOn.coordChange := ContMDiffOn.coordChange -protected theorem Smooth.coordChange (hf : Smooth IM IB f) - (hg : Smooth IM 𝓘(𝕜, F) g) (he : ∀ x, f x ∈ e.baseSet) (he' : ∀ x, f x ∈ e'.baseSet) : - Smooth IM 𝓘(𝕜, F) (fun y ↦ e.coordChange e' (f y) (g y)) := fun x ↦ - (hf x).coordChange (hg x) (he x) (he' x) +@[deprecated (since := "2024-11-21")] +alias Smooth.coordChange := ContMDiff.coordChange variable (e e') @@ -458,8 +432,8 @@ instance SmoothFiberwiseLinear.hasGroupoid : haveI : MemTrivializationAtlas e := ⟨he⟩ haveI : MemTrivializationAtlas e' := ⟨he'⟩ rw [mem_smoothFiberwiseLinear_iff] - refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, smoothOn_coordChangeL e e', - smoothOn_symm_coordChangeL e e', ?_⟩ + refine ⟨_, _, e.open_baseSet.inter e'.open_baseSet, contMDiffOn_coordChangeL e e', + contMDiffOn_symm_coordChangeL e e', ?_⟩ refine PartialHomeomorph.eqOnSourceSetoid.symm ⟨?_, ?_⟩ · simp only [e.symm_trans_source_eq e', FiberwiseLinear.partialHomeomorph, trans_toPartialEquiv, symm_toPartialEquiv] @@ -478,10 +452,10 @@ instance Bundle.TotalSpace.smoothManifoldWithCorners [SmoothManifoldWithCorners refine ⟨ContMDiffOn.congr ?_ (EqOnSource.eqOn heφ), ContMDiffOn.congr ?_ (EqOnSource.eqOn (EqOnSource.symm' heφ))⟩ · rw [EqOnSource.source_eq heφ] - apply smoothOn_fst.prod_mk + apply contMDiffOn_fst.prod_mk exact (hφ.comp contMDiffOn_fst <| prod_subset_preimage_fst _ _).clm_apply contMDiffOn_snd · rw [EqOnSource.target_eq heφ] - apply smoothOn_fst.prod_mk + apply contMDiffOn_fst.prod_mk exact (h2φ.comp contMDiffOn_fst <| prod_subset_preimage_fst _ _).clm_apply contMDiffOn_snd section @@ -517,41 +491,36 @@ theorem Trivialization.contMDiff_iff {f : M → TotalSpace F E} (he : ∀ x, f x ContMDiff IM 𝓘(𝕜, F) n (fun x ↦ (e (f x)).2) := (forall_congr' fun x ↦ e.contMDiffAt_iff (he x)).trans forall_and -theorem Trivialization.smoothWithinAt_iff {f : M → TotalSpace F E} {s : Set M} {x₀ : M} - (he : f x₀ ∈ e.source) : - SmoothWithinAt IM (IB.prod 𝓘(𝕜, F)) f s x₀ ↔ - SmoothWithinAt IM IB (fun x => (f x).proj) s x₀ ∧ - SmoothWithinAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s x₀ := - e.contMDiffWithinAt_iff he +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothWithinAt_iff := Trivialization.contMDiffWithinAt_iff -theorem Trivialization.smoothAt_iff {f : M → TotalSpace F E} {x₀ : M} (he : f x₀ ∈ e.source) : - SmoothAt IM (IB.prod 𝓘(𝕜, F)) f x₀ ↔ - SmoothAt IM IB (fun x => (f x).proj) x₀ ∧ SmoothAt IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) x₀ := - e.contMDiffAt_iff he +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothAt_iff := Trivialization.contMDiffAt_iff -theorem Trivialization.smoothOn_iff {f : M → TotalSpace F E} {s : Set M} - (he : MapsTo f s e.source) : - SmoothOn IM (IB.prod 𝓘(𝕜, F)) f s ↔ - SmoothOn IM IB (fun x => (f x).proj) s ∧ SmoothOn IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) s := - e.contMDiffOn_iff he - -theorem Trivialization.smooth_iff {f : M → TotalSpace F E} (he : ∀ x, f x ∈ e.source) : - Smooth IM (IB.prod 𝓘(𝕜, F)) f ↔ - Smooth IM IB (fun x => (f x).proj) ∧ Smooth IM 𝓘(𝕜, F) (fun x ↦ (e (f x)).2) := - e.contMDiff_iff he - -theorem Trivialization.smoothOn (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : - SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e e.source := by - have : SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) id e.source := smoothOn_id - rw [e.smoothOn_iff (mapsTo_id _)] at this +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothOn_iff := Trivialization.contMDiffOn_iff + +@[deprecated (since := "2024-11-21")] +alias Trivialization.smooth_iff := Trivialization.contMDiff_iff + +theorem Trivialization.contMDiffOn (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : + ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ e e.source := by + have : ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ id e.source := contMDiffOn_id + rw [e.contMDiffOn_iff (mapsTo_id _)] at this exact (this.1.prod_mk this.2).congr fun x hx ↦ (e.mk_proj_snd hx).symm -theorem Trivialization.smoothOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : - SmoothOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) e.toPartialHomeomorph.symm e.target := by - rw [e.smoothOn_iff e.toPartialHomeomorph.symm_mapsTo] - refine ⟨smoothOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, smoothOn_snd.congr fun x hx ↦ ?_⟩ +theorem Trivialization.contMDiffOn_symm (e : Trivialization F (π F E)) [MemTrivializationAtlas e] : + ContMDiffOn (IB.prod 𝓘(𝕜, F)) (IB.prod 𝓘(𝕜, F)) ⊤ e.toPartialHomeomorph.symm e.target := by + rw [e.contMDiffOn_iff e.toPartialHomeomorph.symm_mapsTo] + refine ⟨contMDiffOn_fst.congr fun x hx ↦ e.proj_symm_apply hx, + contMDiffOn_snd.congr fun x hx ↦ ?_⟩ rw [e.apply_symm_apply hx] +@[deprecated (since := "2024-11-21")] alias Trivialization.smoothOn := Trivialization.contMDiffOn + +@[deprecated (since := "2024-11-21")] +alias Trivialization.smoothOn_symm := Trivialization.contMDiffOn_symm + end /-! ### Core construction for smooth vector bundles -/ @@ -563,21 +532,24 @@ variable {ι : Type*} (Z : VectorBundleCore 𝕜 B F ι) /-- Mixin for a `VectorBundleCore` stating smoothness (of transition functions). -/ class IsSmooth (IB : ModelWithCorners 𝕜 EB HB) : Prop where - smoothOn_coordChange : - ∀ i j, SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) + contMDiffOn_coordChange : + ∀ i j, ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) -theorem smoothOn_coordChange (IB : ModelWithCorners 𝕜 EB HB) [h : Z.IsSmooth IB] (i j : ι) : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) := +theorem contMDiffOn_coordChange (IB : ModelWithCorners 𝕜 EB HB) [h : Z.IsSmooth IB] (i j : ι) : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (Z.coordChange i j) (Z.baseSet i ∩ Z.baseSet j) := h.1 i j +@[deprecated (since := "2024-11-21")] +alias smoothOn_coordChange := contMDiffOn_coordChange + variable [Z.IsSmooth IB] /-- If a `VectorBundleCore` has the `IsSmooth` mixin, then the vector bundle constructed from it is a smooth vector bundle. -/ instance smoothVectorBundle : SmoothVectorBundle F Z.Fiber IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro - - ⟨i, rfl⟩ ⟨i', rfl⟩ - refine (Z.smoothOn_coordChange IB i i').congr fun b hb ↦ ?_ + refine (Z.contMDiffOn_coordChange IB i i').congr fun b hb ↦ ?_ ext v exact Z.localTriv_coordChange_eq i i' hb v @@ -587,12 +559,12 @@ end VectorBundleCore /-- A trivial vector bundle over a smooth manifold is a smooth vector bundle. -/ instance Bundle.Trivial.smoothVectorBundle : SmoothVectorBundle F (Bundle.Trivial B F) IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by intro e e' he he' obtain rfl := Bundle.Trivial.eq_trivialization B F e obtain rfl := Bundle.Trivial.eq_trivialization B F e' simp_rw [Bundle.Trivial.trivialization.coordChangeL] - exact smooth_const.smoothOn + exact contMDiff_const.contMDiffOn /-! ### Direct sums of smooth vector bundles -/ @@ -613,15 +585,14 @@ variable [SmoothManifoldWithCorners IB B] /-- The direct sum of two smooth vector bundles over the same base is a smooth vector bundle. -/ instance Bundle.Prod.smoothVectorBundle : SmoothVectorBundle (F₁ × F₂) (E₁ ×ᵇ E₂) IB where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro _ _ ⟨e₁, e₂, i₁, i₂, rfl⟩ ⟨e₁', e₂', i₁', i₂', rfl⟩ - rw [SmoothOn] refine ContMDiffOn.congr ?_ (e₁.coordChangeL_prod 𝕜 e₁' e₂ e₂') refine ContMDiffOn.clm_prodMap ?_ ?_ - · refine (smoothOn_coordChangeL e₁ e₁').mono ?_ + · refine (contMDiffOn_coordChangeL e₁ e₁').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac - · refine (smoothOn_coordChangeL e₂ e₂').mono ?_ + · refine (contMDiffOn_coordChangeL e₂ e₂').mono ?_ simp only [Trivialization.baseSet_prod, mfld_simps] mfld_set_tac @@ -641,7 +612,7 @@ class IsSmooth (a : VectorPrebundle 𝕜 F E) : Prop where exists_smoothCoordChange : ∀ᵉ (e ∈ a.pretrivializationAtlas) (e' ∈ a.pretrivializationAtlas), ∃ f : B → F →L[𝕜] F, - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) f (e.baseSet ∩ e'.baseSet) ∧ + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ f (e.baseSet ∩ e'.baseSet) ∧ ∀ (b : B) (_ : b ∈ e.baseSet ∩ e'.baseSet) (v : F), f b v = (e' ⟨b, e.symm b v⟩).2 @@ -655,11 +626,14 @@ noncomputable def smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) (b : B) : F →L[𝕜] F := Classical.choose (ha.exists_smoothCoordChange e he e' he') b -theorem smoothOn_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) +theorem contMDiffOn_smoothCoordChange (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (a.smoothCoordChange IB he he') (e.baseSet ∩ e'.baseSet) := + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (a.smoothCoordChange IB he he') (e.baseSet ∩ e'.baseSet) := (Classical.choose_spec (ha.exists_smoothCoordChange e he e' he')).1 +@[deprecated (since := "2024-11-21")] +alias smoothOn_smoothCoordChange := contMDiffOn_smoothCoordChange + theorem smoothCoordChange_apply (he : e ∈ a.pretrivializationAtlas) (he' : e' ∈ a.pretrivializationAtlas) {b : B} (hb : b ∈ e.baseSet ∩ e'.baseSet) (v : F) : a.smoothCoordChange IB he he' b v = (e' ⟨b, e.symm b v⟩).2 := @@ -678,9 +652,9 @@ variable (IB) in theorem smoothVectorBundle : @SmoothVectorBundle _ _ F E _ _ _ _ _ _ IB _ _ _ _ _ _ a.totalSpaceTopology _ a.toFiberBundle a.toVectorBundle := letI := a.totalSpaceTopology; letI := a.toFiberBundle; letI := a.toVectorBundle - { smoothOn_coordChangeL := by + { contMDiffOn_coordChangeL := by rintro _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩ - refine (a.smoothOn_smoothCoordChange he he').congr ?_ + refine (a.contMDiffOn_smoothCoordChange he he').congr ?_ intro b hb ext v rw [a.smoothCoordChange_apply he he' hb v, ContinuousLinearEquiv.coe_coe, diff --git a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean index 2a0e077c860aa..0bc3d2285bba8 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/FiberwiseLinear.lean @@ -103,14 +103,14 @@ fiberwise linear partial homeomorphism. -/ theorem SmoothFiberwiseLinear.locality_aux₁ (e : PartialHomeomorph (B × F) (B × F)) (h : ∀ p ∈ e.source, ∃ s : Set (B × F), IsOpen s ∧ p ∈ s ∧ ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr s).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn)) : ∃ U : Set B, e.source = U ×ˢ univ ∧ ∀ x ∈ U, ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (_huU : u ⊆ U) (_hux : x ∈ u), - ∃ (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + ∃ (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr (u ×ˢ univ)).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn) := by rw [SetCoe.forall'] at h @@ -155,13 +155,13 @@ theorem SmoothFiberwiseLinear.locality_aux₂ (e : PartialHomeomorph (B × F) (B (hU : e.source = U ×ˢ univ) (h : ∀ x ∈ U, ∃ (φ : B → F ≃L[𝕜] F) (u : Set B) (hu : IsOpen u) (_hUu : u ⊆ U) (_hux : x ∈ u) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x : F →L[𝕜] F)) u) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((φ x).symm : F →L[𝕜] F)) u), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x : F →L[𝕜] F)) u) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((φ x).symm : F →L[𝕜] F)) u), (e.restr (u ×ˢ univ)).EqOnSource (FiberwiseLinear.partialHomeomorph φ hu hφ.continuousOn h2φ.continuousOn)) : ∃ (Φ : B → F ≃L[𝕜] F) (U : Set B) (hU₀ : IsOpen U) (hΦ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (Φ x : F →L[𝕜] F)) U) (h2Φ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => ((Φ x).symm : F →L[𝕜] F)) U), + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (Φ x : F →L[𝕜] F)) U) (h2Φ : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => ((Φ x).symm : F →L[𝕜] F)) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph Φ hU₀ hΦ.continuousOn h2Φ.continuousOn) := by classical rw [SetCoe.forall'] at h @@ -192,14 +192,14 @@ theorem SmoothFiberwiseLinear.locality_aux₂ (e : PartialHomeomorph (B × F) (B intro x y hyu refine (hΦ y (hUu x hyu)).trans ?_ exact iUnionLift_mk ⟨y, hyu⟩ _ - have hΦ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun y => (Φ y : F →L[𝕜] F)) U := by + have hΦ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun y => (Φ y : F →L[𝕜] F)) U := by apply contMDiffOn_of_locally_contMDiffOn intro x hx refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, ?_⟩ refine (ContMDiffOn.congr (hφ ⟨x, hx⟩) ?_).mono inter_subset_right intro y hy rw [hΦφ ⟨x, hx⟩ y hy] - have h2Φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun y => ((Φ y).symm : F →L[𝕜] F)) U := by + have h2Φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun y => ((Φ y).symm : F →L[𝕜] F)) U := by apply contMDiffOn_of_locally_contMDiffOn intro x hx refine ⟨u ⟨x, hx⟩, hu ⟨x, hx⟩, hux _, ?_⟩ @@ -221,13 +221,13 @@ variable {F B IB} in -- in `smoothFiberwiseLinear` are quite slow, even with this change) private theorem mem_aux {e : PartialHomeomorph (B × F) (B × F)} : (e ∈ ⋃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), {e | e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn)}) ↔ ∃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn) := by simp only [mem_iUnion, mem_setOf_eq] @@ -239,8 +239,8 @@ to the vector bundle belong to this groupoid. -/ def smoothFiberwiseLinear : StructureGroupoid (B × F) where members := ⋃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) - (hφ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) - (h2φ : SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + (hφ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) + (h2φ : ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), {e | e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn)} trans' := by simp only [mem_aux] @@ -248,11 +248,11 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where refine ⟨fun b => (φ b).trans (φ' b), _, hU.inter hU', ?_, ?_, Setoid.trans (PartialHomeomorph.EqOnSource.trans' heφ heφ') ⟨?_, ?_⟩⟩ · show - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x : B => (φ' x).toContinuousLinearMap ∘L (φ x).toContinuousLinearMap) (U ∩ U') exact (hφ'.mono inter_subset_right).clm_comp (hφ.mono inter_subset_left) · show - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x : B => (φ x).symm.toContinuousLinearMap ∘L (φ' x).symm.toContinuousLinearMap) (U ∩ U') exact (h2φ.mono inter_subset_left).clm_comp (h2φ'.mono inter_subset_right) @@ -267,8 +267,8 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where exact hφ id_mem' := by simp_rw [mem_aux] - refine ⟨fun _ ↦ ContinuousLinearEquiv.refl 𝕜 F, univ, isOpen_univ, smoothOn_const, - smoothOn_const, ⟨?_, fun b _hb ↦ rfl⟩⟩ + refine ⟨fun _ ↦ ContinuousLinearEquiv.refl 𝕜 F, univ, isOpen_univ, contMDiffOn_const, + contMDiffOn_const, ⟨?_, fun b _hb ↦ rfl⟩⟩ simp only [FiberwiseLinear.partialHomeomorph, PartialHomeomorph.refl_partialEquiv, PartialEquiv.refl_source, univ_prod_univ] locality' := by @@ -286,7 +286,7 @@ def smoothFiberwiseLinear : StructureGroupoid (B × F) where theorem mem_smoothFiberwiseLinear_iff (e : PartialHomeomorph (B × F) (B × F)) : e ∈ smoothFiberwiseLinear B F IB ↔ ∃ (φ : B → F ≃L[𝕜] F) (U : Set B) (hU : IsOpen U) (hφ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : - SmoothOn IB 𝓘(𝕜, F →L[𝕜] F) (fun x => (φ x).symm : B → F →L[𝕜] F) U), + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => φ x : B → F →L[𝕜] F) U) (h2φ : + ContMDiffOn IB 𝓘(𝕜, F →L[𝕜] F) ⊤ (fun x => (φ x).symm : B → F →L[𝕜] F) U), e.EqOnSource (FiberwiseLinear.partialHomeomorph φ hU hφ.continuousOn h2φ.continuousOn) := mem_aux diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean index f56284f714605..d00cefedfa42b 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Hom.lean @@ -41,16 +41,19 @@ variable {𝕜 B F₁ F₂ M : Type*} {E₁ : B → Type*} {E₂ : B → Type*} local notation "LE₁E₂" => TotalSpace (F₁ →L[𝕜] F₂) (Bundle.ContinuousLinearMap (RingHom.id 𝕜) E₁ E₂) -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11083): moved slow parts to separate lemmas -theorem smoothOn_continuousLinearMapCoordChange +theorem contMDiffOn_continuousLinearMapCoordChange [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle F₂ E₂ IB] [MemTrivializationAtlas e₁] [MemTrivializationAtlas e₁'] [MemTrivializationAtlas e₂] [MemTrivializationAtlas e₂'] : - SmoothOn IB 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₂) + ContMDiffOn IB 𝓘(𝕜, (F₁ →L[𝕜] F₂) →L[𝕜] F₁ →L[𝕜] F₂) ⊤ (continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂') (e₁.baseSet ∩ e₂.baseSet ∩ (e₁'.baseSet ∩ e₂'.baseSet)) := by - have h₁ := smoothOn_coordChangeL (IB := IB) e₁' e₁ - have h₂ := smoothOn_coordChangeL (IB := IB) e₂ e₂' + have h₁ := contMDiffOn_coordChangeL (IB := IB) e₁' e₁ (n := ⊤) + have h₂ := contMDiffOn_coordChangeL (IB := IB) e₂ e₂' (n := ⊤) refine (h₁.mono ?_).cle_arrowCongr (h₂.mono ?_) <;> mfld_set_tac +@[deprecated (since := "2024-11-21")] +alias smoothOn_continuousLinearMapCoordChange := contMDiffOn_continuousLinearMapCoordChange + variable [∀ x, TopologicalAddGroup (E₂ x)] [∀ x, ContinuousSMul 𝕜 (E₂ x)] theorem hom_chart (y₀ y : LE₁E₂) : @@ -67,12 +70,8 @@ theorem contMDiffAt_hom_bundle (f : M → LE₁E₂) {x₀ : M} {n : ℕ∞} : (fun x => inCoordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := contMDiffAt_totalSpace .. -theorem smoothAt_hom_bundle (f : M → LE₁E₂) {x₀ : M} : - SmoothAt IM (IB.prod 𝓘(𝕜, F₁ →L[𝕜] F₂)) f x₀ ↔ - SmoothAt IM IB (fun x => (f x).1) x₀ ∧ - SmoothAt IM 𝓘(𝕜, F₁ →L[𝕜] F₂) - (fun x => inCoordinates F₁ E₁ F₂ E₂ (f x₀).1 (f x).1 (f x₀).1 (f x).1 (f x).2) x₀ := - contMDiffAt_hom_bundle f +@[deprecated (since := "2024-11-21")] alias smoothAt_hom_bundle := contMDiffAt_hom_bundle + variable [SmoothVectorBundle F₁ E₁ IB] [SmoothVectorBundle F₂ E₂ IB] @@ -81,7 +80,7 @@ instance Bundle.ContinuousLinearMap.vectorPrebundle.isSmooth : exists_smoothCoordChange := by rintro _ ⟨e₁, e₂, he₁, he₂, rfl⟩ _ ⟨e₁', e₂', he₁', he₂', rfl⟩ exact ⟨continuousLinearMapCoordChange (RingHom.id 𝕜) e₁ e₁' e₂ e₂', - smoothOn_continuousLinearMapCoordChange, + contMDiffOn_continuousLinearMapCoordChange, continuousLinearMapCoordChange_apply (RingHom.id 𝕜) e₁ e₁' e₂ e₂'⟩ instance SmoothVectorBundle.continuousLinearMap : diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean index ef7514b05054c..7ba8cdb7b86d3 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Pullback.lean @@ -28,14 +28,14 @@ variable [NontriviallyNormedField 𝕜] [∀ x, AddCommMonoid (E x)] [∀ x, Mod [ChartedSpace HB B] {EB' : Type*} [NormedAddCommGroup EB'] [NormedSpace 𝕜 EB'] {HB' : Type*} [TopologicalSpace HB'] (IB' : ModelWithCorners 𝕜 EB' HB') [TopologicalSpace B'] [ChartedSpace HB' B'] [FiberBundle F E] - [VectorBundle 𝕜 F E] [SmoothVectorBundle F E IB] (f : SmoothMap IB' IB B' B) + [VectorBundle 𝕜 F E] [SmoothVectorBundle F E IB] (f : ContMDiffMap IB' IB B' B ⊤) /-- For a smooth vector bundle `E` over a manifold `B` and a smooth map `f : B' → B`, the pullback vector bundle `f *ᵖ E` is a smooth vector bundle. -/ instance SmoothVectorBundle.pullback : SmoothVectorBundle F (f *ᵖ E) IB' where - smoothOn_coordChangeL := by + contMDiffOn_coordChangeL := by rintro _ _ ⟨e, he, rfl⟩ ⟨e', he', rfl⟩ - refine ((smoothOn_coordChangeL e e').comp f.smooth.smoothOn fun b hb => hb).congr ?_ + refine ((contMDiffOn_coordChangeL e e').comp f.contMDiff.contMDiffOn fun b hb => hb).congr ?_ rintro b (hb : f b ∈ e.baseSet ∩ e'.baseSet); ext v show ((e.pullback f).coordChangeL 𝕜 (e'.pullback f) b) v = (e.coordChangeL 𝕜 e' (f b)) v rw [e.coordChangeL_apply e' hb, (e.pullback f).coordChangeL_apply' _] diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 198e5170ad73e..751bf15036342 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -39,9 +39,7 @@ structure ContMDiffSection where protected contMDiff_toFun : ContMDiff I (I.prod 𝓘(𝕜, F)) n fun x ↦ TotalSpace.mk' F x (toFun x) -/-- Bundled smooth sections of a vector bundle. -/ -abbrev SmoothSection := - ContMDiffSection I F ⊤ V +@[deprecated (since := "024-11-21")] alias SmoothSection := ContMDiffSection @[inherit_doc] scoped[Manifold] notation "Cₛ^" n "⟮" I "; " F ", " V "⟯" => ContMDiffSection I F n V @@ -65,9 +63,7 @@ protected theorem contMDiff (s : Cₛ^n⟮I; F, V⟯) : ContMDiff I (I.prod 𝓘(𝕜, F)) n fun x => TotalSpace.mk' F x (s x : V x) := s.contMDiff_toFun -protected theorem smooth (s : Cₛ^∞⟮I; F, V⟯) : - Smooth I (I.prod 𝓘(𝕜, F)) fun x => TotalSpace.mk' F x (s x : V x) := - s.contMDiff_toFun +@[deprecated (since := "2024-11-21")] alias smooth := ContMDiffSection.contMDiff theorem coe_inj ⦃s t : Cₛ^n⟮I; F, V⟯⦄ (h : (s : ∀ x, V x) = t) : s = t := DFunLike.ext' h @@ -114,7 +110,7 @@ theorem coe_sub (s t : Cₛ^n⟮I; F, V⟯) : ⇑(s - t) = s - t := rfl instance instZero : Zero Cₛ^n⟮I; F, V⟯ := - ⟨⟨fun _ => 0, (smooth_zeroSection 𝕜 V).of_le le_top⟩⟩ + ⟨⟨fun _ => 0, (contMDiff_zeroSection 𝕜 V).of_le le_top⟩⟩ instance inhabited : Inhabited Cₛ^n⟮I; F, V⟯ := ⟨0⟩ diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index fd5bda4a24a80..f40f1f950c07a 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -286,7 +286,7 @@ end TangentBundle instance tangentBundleCore.isSmooth : (tangentBundleCore I M).IsSmooth I := by refine ⟨fun i j => ?_⟩ - rw [SmoothOn, contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas i.2), + rw [contMDiffOn_iff_source_of_mem_maximalAtlas (subset_maximalAtlas i.2), contMDiffOn_iff_contDiffOn] · refine ((contDiffOn_fderiv_coord_change (I := I) i j).congr fun x hx => ?_).mono ?_ · rw [PartialEquiv.trans_source'] at hx diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 2c5558360b39f..5551413ea3416 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -52,7 +52,7 @@ def embeddingPiTangent : C^∞⟮I, M; 𝓘(ℝ, ι → E × ℝ), ι → E × val x i := (f i x • extChartAt I (f.c i) x, f i x) property := contMDiff_pi_space.2 fun i => - ((f i).smooth_smul contMDiffOn_extChartAt).prod_mk_space (f i).smooth + ((f i).contMDiff_smul contMDiffOn_extChartAt).prod_mk_space (f i).contMDiff @[local simp] theorem embeddingPiTangent_coe : @@ -81,7 +81,8 @@ theorem comp_embeddingPiTangent_mfderiv (x : M) (hx : x ∈ s) : set L := (ContinuousLinearMap.fst ℝ E ℝ).comp (@ContinuousLinearMap.proj ℝ _ ι (fun _ => E × ℝ) _ _ (fun _ => inferInstance) (f.ind x hx)) - have := L.hasMFDerivAt.comp x f.embeddingPiTangent.smooth.mdifferentiableAt.hasMFDerivAt + have := L.hasMFDerivAt.comp x + (f.embeddingPiTangent.contMDiff.mdifferentiableAt le_top).hasMFDerivAt convert hasMFDerivAt_unique this _ refine (hasMFDerivAt_extChartAt (f.mem_chartAt_ind_source x hx)).congr_of_eventuallyEq ?_ refine (f.eventuallyEq_one x hx).mono fun y hy => ?_ @@ -106,7 +107,7 @@ supports of bump functions, then for some `n` it can be immersed into the `n`-di Euclidean space. -/ theorem exists_immersion_euclidean {ι : Type*} [Finite ι] (f : SmoothBumpCovering ι I M) : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ Injective e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + ContMDiff I (𝓡 n) ⊤ e ∧ Injective e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by cases nonempty_fintype ι set F := EuclideanSpace ℝ (Fin <| finrank ℝ (ι → E × ℝ)) letI : IsNoetherian ℝ (E × ℝ) := IsNoetherian.iff_fg.2 inferInstance @@ -114,10 +115,10 @@ theorem exists_immersion_euclidean {ι : Type*} [Finite ι] (f : SmoothBumpCover set eEF : (ι → E × ℝ) ≃L[ℝ] F := ContinuousLinearEquiv.ofFinrankEq finrank_euclideanSpace_fin.symm refine ⟨_, eEF ∘ f.embeddingPiTangent, - eEF.toDiffeomorph.smooth.comp f.embeddingPiTangent.smooth, + eEF.toDiffeomorph.contMDiff.comp f.embeddingPiTangent.contMDiff, eEF.injective.comp f.embeddingPiTangent_injective, fun x => ?_⟩ rw [mfderiv_comp _ eEF.differentiableAt.mdifferentiableAt - f.embeddingPiTangent.smooth.mdifferentiableAt, + (f.embeddingPiTangent.contMDiff.mdifferentiableAt le_top), eEF.mfderiv_eq] exact eEF.injective.comp (f.embeddingPiTangent_injective_mfderiv _ trivial) @@ -128,7 +129,7 @@ supports of bump functions, then for some `n` it can be embedded into the `n`-di Euclidean space. -/ theorem exists_embedding_euclidean_of_compact [T2Space M] [CompactSpace M] : ∃ (n : ℕ) (e : M → EuclideanSpace ℝ (Fin n)), - Smooth I (𝓡 n) e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by + ContMDiff I (𝓡 n) ⊤ e ∧ IsClosedEmbedding e ∧ ∀ x : M, Injective (mfderiv I (𝓡 n) e x) := by rcases SmoothBumpCovering.exists_isSubordinate I isClosed_univ fun (x : M) _ => univ_mem with ⟨ι, f, -⟩ haveI := f.fintype From f801de4934457721044d87bd482f67a3219c30c1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 09:33:04 +0000 Subject: [PATCH 127/250] chore: split notation typeclasses out of Algebra.Group.Defs (#19350) --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Defs.lean | 192 +---------------------- Mathlib/Algebra/Group/Operations.lean | 215 ++++++++++++++++++++++++++ Mathlib/Data/Part.lean | 2 +- 4 files changed, 219 insertions(+), 191 deletions(-) create mode 100644 Mathlib/Algebra/Group/Operations.lean diff --git a/Mathlib.lean b/Mathlib.lean index 56e4a03916d8e..4b515136d2aa3 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -287,6 +287,7 @@ import Mathlib.Algebra.Group.Nat.Even import Mathlib.Algebra.Group.Nat.TypeTags import Mathlib.Algebra.Group.Nat.Units import Mathlib.Algebra.Group.NatPowAssoc +import Mathlib.Algebra.Group.Operations import Mathlib.Algebra.Group.Opposite import Mathlib.Algebra.Group.PNatPowAssoc import Mathlib.Algebra.Group.Pi.Basic diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 3abfe30a7d417..3d069aab663a3 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -6,12 +6,9 @@ Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro import Mathlib.Data.Int.Notation import Mathlib.Data.Nat.BinaryRec import Mathlib.Algebra.Group.ZeroOne +import Mathlib.Algebra.Group.Operations import Mathlib.Logic.Function.Defs -import Mathlib.Tactic.Lemma -import Mathlib.Tactic.TypeStar import Mathlib.Tactic.Simps.Basic -import Mathlib.Tactic.ToAdditive -import Mathlib.Util.AssertExists import Batteries.Logic /-! @@ -29,21 +26,15 @@ The file does not contain any lemmas except for For basic lemmas about these classes see `Algebra.Group.Basic`. -We also introduce notation classes `SMul` and `VAdd` for multiplicative and additive -actions and register the following instances: +We register the following instances: - `Pow M ℕ`, for monoids `M`, and `Pow G ℤ` for groups `G`; - `SMul ℕ M` for additive monoids `M`, and `SMul ℤ G` for additive groups `G`. -`SMul` is typically, but not exclusively, used for scalar multiplication-like operators. -See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition). - ## Notation - `+`, `-`, `*`, `/`, `^` : the usual arithmetic operations; the underlying functions are `Add.add`, `Neg.neg`/`Sub.sub`, `Mul.mul`, `Div.div`, and `HPow.hPow`. -- `a • b` is used as notation for `HSMul.hSMul a b`. -- `a +ᵥ b` is used as notation for `HVAdd.hVAdd a b`. -/ @@ -55,187 +46,8 @@ universe u v w open Function -/-- -The notation typeclass for heterogeneous additive actions. -This enables the notation `a +ᵥ b : γ` where `a : α`, `b : β`. --/ -class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where - /-- `a +ᵥ b` computes the sum of `a` and `b`. - The meaning of this notation is type-dependent. -/ - hVAdd : α → β → γ - -/-- -The notation typeclass for heterogeneous scalar multiplication. -This enables the notation `a • b : γ` where `a : α`, `b : β`. - -It is assumed to represent a left action in some sense. -The notation `a • b` is augmented with a macro (below) to have it elaborate as a left action. -Only the `b` argument participates in the elaboration algorithm: the algorithm uses the type of `b` -when calculating the type of the surrounding arithmetic expression -and it tries to insert coercions into `b` to get some `b'` -such that `a • b'` has the same type as `b'`. -See the module documentation near the macro for more details. --/ -class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where - /-- `a • b` computes the product of `a` and `b`. - The meaning of this notation is type-dependent, but it is intended to be used for left actions. -/ - hSMul : α → β → γ - -attribute [notation_class smul Simps.copySecond] HSMul -attribute [notation_class nsmul Simps.nsmulArgs] HSMul -attribute [notation_class zsmul Simps.zsmulArgs] HSMul - -/-- Type class for the `+ᵥ` notation. -/ -class VAdd (G : Type u) (P : Type v) where - /-- `a +ᵥ b` computes the sum of `a` and `b`. The meaning of this notation is type-dependent, - but it is intended to be used for left actions. -/ - vadd : G → P → P - -/-- Type class for the `-ᵥ` notation. -/ -class VSub (G : outParam Type*) (P : Type*) where - /-- `a -ᵥ b` computes the difference of `a` and `b`. The meaning of this notation is - type-dependent, but it is intended to be used for additive torsors. -/ - vsub : P → P → G - -/-- Typeclass for types with a scalar multiplication operation, denoted `•` (`\bu`) -/ -@[to_additive (attr := ext)] -class SMul (M : Type u) (α : Type v) where - /-- `a • b` computes the product of `a` and `b`. The meaning of this notation is type-dependent, - but it is intended to be used for left actions. -/ - smul : M → α → α - -@[inherit_doc] infixl:65 " +ᵥ " => HVAdd.hVAdd -@[inherit_doc] infixl:65 " -ᵥ " => VSub.vsub -@[inherit_doc] infixr:73 " • " => HSMul.hSMul - -/-! -We have a macro to make `x • y` notation participate in the expression tree elaborator, -like other arithmetic expressions such as `+`, `*`, `/`, `^`, `=`, inequalities, etc. -The macro is using the `leftact%` elaborator introduced in -[this RFC](https://github.com/leanprover/lean4/issues/2854). - -As a concrete example of the effect of this macro, consider -```lean -variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N) -#check m + r • n -``` -Without the macro, the expression would elaborate as `m + ↑(r • n : ↑N) : M`. -With the macro, the expression elaborates as `m + r • (↑n : M) : M`. -To get the first interpretation, one can write `m + (r • n :)`. - -Here is a quick review of the expression tree elaborator: -1. It builds up an expression tree of all the immediately accessible operations - that are marked with `binop%`, `unop%`, `leftact%`, `rightact%`, `binrel%`, etc. -2. It elaborates every leaf term of this tree - (without an expected type, so as if it were temporarily wrapped in `(... :)`). -3. Using the types of each elaborated leaf, it computes a supremum type they can all be - coerced to, if such a supremum exists. -4. It inserts coercions around leaf terms wherever needed. - -The hypothesis is that individual expression trees tend to be calculations with respect -to a single algebraic structure. - -Note(kmill): If we were to remove `HSMul` and switch to using `SMul` directly, -then the expression tree elaborator would not be able to insert coercions within the right operand; -they would likely appear as `↑(x • y)` rather than `x • ↑y`, unlike other arithmetic operations. --/ - -@[inherit_doc HSMul.hSMul] -macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) - -attribute [to_additive existing] Mul Div HMul instHMul HDiv instHDiv HSMul -attribute [to_additive (reorder := 1 2) SMul] Pow -attribute [to_additive (reorder := 1 2)] HPow -attribute [to_additive existing (reorder := 1 2, 5 6) hSMul] HPow.hPow -attribute [to_additive existing (reorder := 1 2, 4 5) smul] Pow.pow - -@[to_additive (attr := default_instance)] -instance instHSMul {α β} [SMul α β] : HSMul α β β where - hSMul := SMul.smul - -@[to_additive] -theorem SMul.smul_eq_hSMul {α β} [SMul α β] : (SMul.smul : α → β → β) = HSMul.hSMul := rfl - -attribute [to_additive existing (reorder := 1 2)] instHPow - variable {G : Type*} -/-- Class of types that have an inversion operation. -/ -@[to_additive, notation_class] -class Inv (α : Type u) where - /-- Invert an element of α. -/ - inv : α → α - -@[inherit_doc] -postfix:max "⁻¹" => Inv.inv - -section ite -variable {α : Type*} (P : Prop) [Decidable P] - -section Mul -variable [Mul α] - -@[to_additive] -lemma mul_dite (a : α) (b : P → α) (c : ¬ P → α) : - (a * if h : P then b h else c h) = if h : P then a * b h else a * c h := by split <;> rfl - -@[to_additive] -lemma mul_ite (a b c : α) : (a * if P then b else c) = if P then a * b else a * c := mul_dite .. - -@[to_additive] -lemma dite_mul (a : P → α) (b : ¬ P → α) (c : α) : - (if h : P then a h else b h) * c = if h : P then a h * c else b h * c := by split <;> rfl - -@[to_additive] -lemma ite_mul (a b c : α) : (if P then a else b) * c = if P then a * c else b * c := dite_mul .. - --- We make `mul_ite` and `ite_mul` simp lemmas, but not `add_ite` or `ite_add`. --- The problem we're trying to avoid is dealing with sums of the form `∑ x ∈ s, (f x + ite P 1 0)`, --- in which `add_ite` followed by `sum_ite` would needlessly slice up --- the `f x` terms according to whether `P` holds at `x`. --- There doesn't appear to be a corresponding difficulty so far with `mul_ite` and `ite_mul`. -attribute [simp] mul_dite dite_mul mul_ite ite_mul - -@[to_additive] -lemma dite_mul_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : - ((if h : P then a h else b h) * if h : P then c h else d h) = - if h : P then a h * c h else b h * d h := by split <;> rfl - -@[to_additive] -lemma ite_mul_ite (a b c d : α) : - ((if P then a else b) * if P then c else d) = if P then a * c else b * d := by split <;> rfl - -end Mul - -section Div -variable [Div α] - -@[to_additive] -lemma div_dite (a : α) (b : P → α) (c : ¬ P → α) : - (a / if h : P then b h else c h) = if h : P then a / b h else a / c h := by split <;> rfl - -@[to_additive] -lemma div_ite (a b c : α) : (a / if P then b else c) = if P then a / b else a / c := div_dite .. - -@[to_additive] -lemma dite_div (a : P → α) (b : ¬ P → α) (c : α) : - (if h : P then a h else b h) / c = if h : P then a h / c else b h / c := by split <;> rfl - -@[to_additive] -lemma ite_div (a b c : α) : (if P then a else b) / c = if P then a / c else b / c := dite_div .. - -@[to_additive] -lemma dite_div_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : - ((if h : P then a h else b h) / if h : P then c h else d h) = - if h : P then a h / c h else b h / d h := by split <;> rfl - -@[to_additive] -lemma ite_div_ite (a b c d : α) : - ((if P then a else b) / if P then c else d) = if P then a / c else b / d := dite_div_dite .. - -end Div -end ite - section Mul variable [Mul G] diff --git a/Mathlib/Algebra/Group/Operations.lean b/Mathlib/Algebra/Group/Operations.lean new file mode 100644 index 0000000000000..dcc6957a258a7 --- /dev/null +++ b/Mathlib/Algebra/Group/Operations.lean @@ -0,0 +1,215 @@ +/- +Copyright (c) 2014 Jeremy Avigad. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jeremy Avigad, Leonardo de Moura, Simon Hudon, Mario Carneiro +-/ + +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.TypeStar +import Mathlib.Tactic.ToAdditive +import Mathlib.Util.AssertExists + +/-! +# Typeclasses for algebraic operations + +Notation typeclass for `Inv`, the multiplicative analogue of `Neg`. + +We also introduce notation classes `SMul` and `VAdd` for multiplicative and additive +actions. + +`SMul` is typically, but not exclusively, used for scalar multiplication-like operators. +See the module `Algebra.AddTorsor` for a motivating example for the name `VAdd` (vector addition). + +## Notation + +- `a • b` is used as notation for `HSMul.hSMul a b`. +- `a +ᵥ b` is used as notation for `HVAdd.hVAdd a b`. + +-/ + +assert_not_exists One +assert_not_exists Function.Injective + +universe u v w + + +/-- +The notation typeclass for heterogeneous additive actions. +This enables the notation `a +ᵥ b : γ` where `a : α`, `b : β`. +-/ +class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where + /-- `a +ᵥ b` computes the sum of `a` and `b`. + The meaning of this notation is type-dependent. -/ + hVAdd : α → β → γ + +/-- +The notation typeclass for heterogeneous scalar multiplication. +This enables the notation `a • b : γ` where `a : α`, `b : β`. + +It is assumed to represent a left action in some sense. +The notation `a • b` is augmented with a macro (below) to have it elaborate as a left action. +Only the `b` argument participates in the elaboration algorithm: the algorithm uses the type of `b` +when calculating the type of the surrounding arithmetic expression +and it tries to insert coercions into `b` to get some `b'` +such that `a • b'` has the same type as `b'`. +See the module documentation near the macro for more details. +-/ +class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where + /-- `a • b` computes the product of `a` and `b`. + The meaning of this notation is type-dependent, but it is intended to be used for left actions. -/ + hSMul : α → β → γ + +attribute [notation_class smul Simps.copySecond] HSMul +attribute [notation_class nsmul Simps.nsmulArgs] HSMul +attribute [notation_class zsmul Simps.zsmulArgs] HSMul + +/-- Type class for the `+ᵥ` notation. -/ +class VAdd (G : Type u) (P : Type v) where + /-- `a +ᵥ b` computes the sum of `a` and `b`. The meaning of this notation is type-dependent, + but it is intended to be used for left actions. -/ + vadd : G → P → P + +/-- Type class for the `-ᵥ` notation. -/ +class VSub (G : outParam Type*) (P : Type*) where + /-- `a -ᵥ b` computes the difference of `a` and `b`. The meaning of this notation is + type-dependent, but it is intended to be used for additive torsors. -/ + vsub : P → P → G + +/-- Typeclass for types with a scalar multiplication operation, denoted `•` (`\bu`) -/ +@[to_additive (attr := ext)] +class SMul (M : Type u) (α : Type v) where + /-- `a • b` computes the product of `a` and `b`. The meaning of this notation is type-dependent, + but it is intended to be used for left actions. -/ + smul : M → α → α + +@[inherit_doc] infixl:65 " +ᵥ " => HVAdd.hVAdd +@[inherit_doc] infixl:65 " -ᵥ " => VSub.vsub +@[inherit_doc] infixr:73 " • " => HSMul.hSMul + +/-! +We have a macro to make `x • y` notation participate in the expression tree elaborator, +like other arithmetic expressions such as `+`, `*`, `/`, `^`, `=`, inequalities, etc. +The macro is using the `leftact%` elaborator introduced in +[this RFC](https://github.com/leanprover/lean4/issues/2854). + +As a concrete example of the effect of this macro, consider +```lean +variable [Ring R] [AddCommMonoid M] [Module R M] (r : R) (N : Submodule R M) (m : M) (n : N) +#check m + r • n +``` +Without the macro, the expression would elaborate as `m + ↑(r • n : ↑N) : M`. +With the macro, the expression elaborates as `m + r • (↑n : M) : M`. +To get the first interpretation, one can write `m + (r • n :)`. + +Here is a quick review of the expression tree elaborator: +1. It builds up an expression tree of all the immediately accessible operations + that are marked with `binop%`, `unop%`, `leftact%`, `rightact%`, `binrel%`, etc. +2. It elaborates every leaf term of this tree + (without an expected type, so as if it were temporarily wrapped in `(... :)`). +3. Using the types of each elaborated leaf, it computes a supremum type they can all be + coerced to, if such a supremum exists. +4. It inserts coercions around leaf terms wherever needed. + +The hypothesis is that individual expression trees tend to be calculations with respect +to a single algebraic structure. + +Note(kmill): If we were to remove `HSMul` and switch to using `SMul` directly, +then the expression tree elaborator would not be able to insert coercions within the right operand; +they would likely appear as `↑(x • y)` rather than `x • ↑y`, unlike other arithmetic operations. +-/ + +@[inherit_doc HSMul.hSMul] +macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) + +attribute [to_additive existing] Mul Div HMul instHMul HDiv instHDiv HSMul +attribute [to_additive (reorder := 1 2) SMul] Pow +attribute [to_additive (reorder := 1 2)] HPow +attribute [to_additive existing (reorder := 1 2, 5 6) hSMul] HPow.hPow +attribute [to_additive existing (reorder := 1 2, 4 5) smul] Pow.pow + +@[to_additive (attr := default_instance)] +instance instHSMul {α β} [SMul α β] : HSMul α β β where + hSMul := SMul.smul + +@[to_additive] +theorem SMul.smul_eq_hSMul {α β} [SMul α β] : (SMul.smul : α → β → β) = HSMul.hSMul := rfl + +attribute [to_additive existing (reorder := 1 2)] instHPow + +variable {G : Type*} + +/-- Class of types that have an inversion operation. -/ +@[to_additive, notation_class] +class Inv (α : Type u) where + /-- Invert an element of α. -/ + inv : α → α + +@[inherit_doc] +postfix:max "⁻¹" => Inv.inv + +section ite +variable {α : Type*} (P : Prop) [Decidable P] + +section Mul +variable [Mul α] + +@[to_additive] +lemma mul_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a * if h : P then b h else c h) = if h : P then a * b h else a * c h := by split <;> rfl + +@[to_additive] +lemma mul_ite (a b c : α) : (a * if P then b else c) = if P then a * b else a * c := mul_dite .. + +@[to_additive] +lemma dite_mul (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) * c = if h : P then a h * c else b h * c := by split <;> rfl + +@[to_additive] +lemma ite_mul (a b c : α) : (if P then a else b) * c = if P then a * c else b * c := dite_mul .. + +-- We make `mul_ite` and `ite_mul` simp lemmas, but not `add_ite` or `ite_add`. +-- The problem we're trying to avoid is dealing with sums of the form `∑ x ∈ s, (f x + ite P 1 0)`, +-- in which `add_ite` followed by `sum_ite` would needlessly slice up +-- the `f x` terms according to whether `P` holds at `x`. +-- There doesn't appear to be a corresponding difficulty so far with `mul_ite` and `ite_mul`. +attribute [simp] mul_dite dite_mul mul_ite ite_mul + +@[to_additive] +lemma dite_mul_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) * if h : P then c h else d h) = + if h : P then a h * c h else b h * d h := by split <;> rfl + +@[to_additive] +lemma ite_mul_ite (a b c d : α) : + ((if P then a else b) * if P then c else d) = if P then a * c else b * d := by split <;> rfl + +end Mul + +section Div +variable [Div α] + +@[to_additive] +lemma div_dite (a : α) (b : P → α) (c : ¬ P → α) : + (a / if h : P then b h else c h) = if h : P then a / b h else a / c h := by split <;> rfl + +@[to_additive] +lemma div_ite (a b c : α) : (a / if P then b else c) = if P then a / b else a / c := div_dite .. + +@[to_additive] +lemma dite_div (a : P → α) (b : ¬ P → α) (c : α) : + (if h : P then a h else b h) / c = if h : P then a h / c else b h / c := by split <;> rfl + +@[to_additive] +lemma ite_div (a b c : α) : (if P then a else b) / c = if P then a / c else b / c := dite_div .. + +@[to_additive] +lemma dite_div_dite (a : P → α) (b : ¬ P → α) (c : P → α) (d : ¬ P → α) : + ((if h : P then a h else b h) / if h : P then c h else d h) = + if h : P then a h / c h else b h / d h := by split <;> rfl + +@[to_additive] +lemma ite_div_ite (a b c d : α) : + ((if P then a else b) / if P then c else d) = if P then a / c else b / d := dite_div_dite .. + +end Div +end ite diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index 58395794466e5..842a941412357 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -5,7 +5,7 @@ Authors: Mario Carneiro, Jeremy Avigad, Simon Hudon -/ import Mathlib.Data.Set.Subsingleton import Mathlib.Logic.Equiv.Defs -import Mathlib.Algebra.Group.Defs +import Mathlib.Algebra.Group.Operations /-! # Partial values of a type From 725d2ded25e22f4f26c9554e47a48398db571c9b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 10:13:39 +0000 Subject: [PATCH 128/250] chore: rename `CompleteLattice.Independent`/`.SetIndependent` to `iSupIndep`/`sSupIndep` (#19409) These names are shorter, closer to `SupIndep` and clearer without their namespace. --- Counterexamples/DirectSumIsInternal.lean | 8 +- Mathlib/Algebra/DirectSum/Basic.lean | 4 +- Mathlib/Algebra/DirectSum/Internal.lean | 2 +- Mathlib/Algebra/DirectSum/LinearMap.lean | 12 +- Mathlib/Algebra/DirectSum/Module.lean | 73 +++--- Mathlib/Algebra/DirectSum/Ring.lean | 5 +- Mathlib/Algebra/Lie/Semisimple/Basic.lean | 6 +- Mathlib/Algebra/Lie/Semisimple/Defs.lean | 2 +- Mathlib/Algebra/Lie/Submodule.lean | 9 +- Mathlib/Algebra/Lie/TraceForm.lean | 8 +- Mathlib/Algebra/Lie/Weights/Basic.lean | 33 ++- Mathlib/Algebra/Lie/Weights/Chain.lean | 6 +- Mathlib/Algebra/Lie/Weights/Killing.lean | 4 +- Mathlib/Algebra/Module/Torsion.lean | 15 +- Mathlib/Analysis/InnerProductSpace/Basic.lean | 4 +- .../InnerProductSpace/Projection.lean | 2 +- Mathlib/GroupTheory/NoncommPiCoprod.lean | 28 ++- Mathlib/GroupTheory/Sylow.lean | 2 +- Mathlib/LinearAlgebra/DFinsupp.lean | 110 +++++---- Mathlib/LinearAlgebra/Dimension/Finite.lean | 21 +- Mathlib/LinearAlgebra/Eigenspace/Basic.lean | 16 +- Mathlib/LinearAlgebra/Eigenspace/Matrix.lean | 2 +- Mathlib/LinearAlgebra/Eigenspace/Pi.lean | 4 +- Mathlib/LinearAlgebra/LinearIndependent.lean | 11 +- .../Projectivization/Independence.lean | 8 +- Mathlib/Order/CompactlyGenerated/Basic.lean | 127 ++++++---- Mathlib/Order/SupIndep.lean | 223 ++++++++++++------ Mathlib/RingTheory/FiniteLength.lean | 8 +- Mathlib/RingTheory/Noetherian/Defs.lean | 13 +- Mathlib/RingTheory/SimpleModule.lean | 10 +- scripts/nolints_prime_decls.txt | 12 +- 31 files changed, 486 insertions(+), 302 deletions(-) diff --git a/Counterexamples/DirectSumIsInternal.lean b/Counterexamples/DirectSumIsInternal.lean index db6389188d856..7632fb6414a31 100644 --- a/Counterexamples/DirectSumIsInternal.lean +++ b/Counterexamples/DirectSumIsInternal.lean @@ -12,10 +12,10 @@ import Mathlib.Tactic.FinCases # Not all complementary decompositions of a module over a semiring make up a direct sum This shows that while `ℤ≤0` and `ℤ≥0` are complementary `ℕ`-submodules of `ℤ`, which in turn -implies as a collection they are `CompleteLattice.Independent` and that they span all of `ℤ`, they +implies as a collection they are `iSupIndep` and that they span all of `ℤ`, they do not form a decomposition into a direct sum. -This file demonstrates why `DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top` must +This file demonstrates why `DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top` must take `Ring R` and not `Semiring R`. -/ @@ -57,9 +57,9 @@ theorem withSign.isCompl : IsCompl ℤ≥0 ℤ≤0 := by · exact Submodule.mem_sup_left (mem_withSign_one.mpr hp) · exact Submodule.mem_sup_right (mem_withSign_neg_one.mpr hn) -def withSign.independent : CompleteLattice.Independent withSign := by +def withSign.independent : iSupIndep withSign := by apply - (CompleteLattice.independent_pair UnitsInt.one_ne_neg_one _).mpr withSign.isCompl.disjoint + (iSupIndep_pair UnitsInt.one_ne_neg_one _).mpr withSign.isCompl.disjoint intro i fin_cases i <;> simp diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index 4aab3891bf560..50045a6c1ba5b 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -367,8 +367,8 @@ theorem coe_of_apply {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S `M` is said to be internal if the canonical map `(⨁ i, A i) →+ M` is bijective. For the alternate statement in terms of independence and spanning, see -`DirectSum.subgroup_isInternal_iff_independent_and_supr_eq_top` and -`DirectSum.isInternal_submodule_iff_independent_and_iSup_eq_top`. -/ +`DirectSum.subgroup_isInternal_iff_iSupIndep_and_supr_eq_top` and +`DirectSum.isInternal_submodule_iff_iSupIndep_and_iSup_eq_top`. -/ def IsInternal {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S M] [AddSubmonoidClass S M] (A : ι → S) : Prop := Function.Bijective (DirectSum.coeAddMonoidHom A) diff --git a/Mathlib/Algebra/DirectSum/Internal.lean b/Mathlib/Algebra/DirectSum/Internal.lean index 34fbe957a6174..a75722faa88dc 100644 --- a/Mathlib/Algebra/DirectSum/Internal.lean +++ b/Mathlib/Algebra/DirectSum/Internal.lean @@ -28,7 +28,7 @@ to represent this case, `(h : DirectSum.IsInternal A) [SetLike.GradedMonoid A]` needed. In the future there will likely be a data-carrying, constructive, typeclass version of `DirectSum.IsInternal` for providing an explicit decomposition function. -When `CompleteLattice.Independent (Set.range A)` (a weaker condition than +When `iSupIndep (Set.range A)` (a weaker condition than `DirectSum.IsInternal A`), these provide a grading of `⨆ i, A i`, and the mapping `⨁ i, A i →+ ⨆ i, A i` can be obtained as `DirectSum.toAddMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. diff --git a/Mathlib/Algebra/DirectSum/LinearMap.lean b/Mathlib/Algebra/DirectSum/LinearMap.lean index 25ffc82193de4..ff8f8b3d653a4 100644 --- a/Mathlib/Algebra/DirectSum/LinearMap.lean +++ b/Mathlib/Algebra/DirectSum/LinearMap.lean @@ -82,8 +82,8 @@ lemma trace_eq_zero_of_mapsTo_ne (h : IsInternal N) [IsNoetherian R M] (σ : ι → ι) (hσ : ∀ i, σ i ≠ i) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N <| σ i)) : trace R M f = 0 := by - have hN : {i | N i ≠ ⊥}.Finite := CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent - h.submodule_independent + have hN : {i | N i ≠ ⊥}.Finite := WellFoundedGT.finite_ne_bot_of_iSupIndep + h.submodule_iSupIndep let s := hN.toFinset let κ := fun i ↦ Module.Free.ChooseBasisIndex R (N i) let b : (i : s) → Basis (κ i) R (N i) := fun i ↦ Module.Free.chooseBasis R (N i) @@ -109,10 +109,10 @@ lemma trace_comp_eq_zero_of_commute_of_trace_restrict_eq_zero (f.mapsTo_maxGenEigenspace_of_comm rfl μ) suffices ∀ μ, trace R _ ((g ∘ₗ f).restrict (hfg μ)) = 0 by classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top f.independent_maxGenEigenspace hf have h_fin : {μ | f.maxGenEigenspace μ ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent f.independent_maxGenEigenspace + WellFoundedGT.finite_ne_bot_of_iSupIndep f.independent_maxGenEigenspace simp [trace_eq_sum_trace_restrict' hds h_fin hfg, this] intro μ replace h_comm : Commute (g.restrict (f.mapsTo_maxGenEigenspace_of_comm h_comm μ)) @@ -136,14 +136,14 @@ Note that it is important the statement gives the user definitional control over _type_ of the term `trace R p (f.restrict hp')` depends on `p`. -/ lemma trace_eq_sum_trace_restrict_of_eq_biSup [∀ i, Module.Finite R (N i)] [∀ i, Module.Free R (N i)] - (s : Finset ι) (h : CompleteLattice.Independent <| fun i : s ↦ N i) + (s : Finset ι) (h : iSupIndep <| fun i : s ↦ N i) {f : Module.End R M} (hf : ∀ i, MapsTo f (N i) (N i)) (p : Submodule R M) (hp : p = ⨆ i ∈ s, N i) (hp' : MapsTo f p p := hp ▸ mapsTo_biSup_of_mapsTo (s : Set ι) hf) : trace R p (f.restrict hp') = ∑ i ∈ s, trace R (N i) (f.restrict (hf i)) := by classical let N' : s → Submodule R p := fun i ↦ (N i).comap p.subtype - replace h : IsInternal N' := hp ▸ isInternal_biSup_submodule_of_independent (s : Set ι) h + replace h : IsInternal N' := hp ▸ isInternal_biSup_submodule_of_iSupIndep (s : Set ι) h have hf' : ∀ i, MapsTo (restrict f hp') (N' i) (N' i) := fun i x hx' ↦ by simpa using hf i hx' let e : (i : s) → N' i ≃ₗ[R] N i := fun ⟨i, hi⟩ ↦ (N i).comapSubtypeEquivOfLe (hp ▸ le_biSup N hi) have _i1 : ∀ i, Module.Finite R (N' i) := fun i ↦ Module.Finite.equiv (e i).symm diff --git a/Mathlib/Algebra/DirectSum/Module.lean b/Mathlib/Algebra/DirectSum/Module.lean index 299c5d4caad54..dba8c5158d5d0 100644 --- a/Mathlib/Algebra/DirectSum/Module.lean +++ b/Mathlib/Algebra/DirectSum/Module.lean @@ -319,8 +319,11 @@ theorem IsInternal.submodule_iSup_eq_top (h : IsInternal A) : iSup A = ⊤ := by exact Function.Bijective.surjective h /-- If a direct sum of submodules is internal then the submodules are independent. -/ -theorem IsInternal.submodule_independent (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_lsum_injective _ h.injective +theorem IsInternal.submodule_iSupIndep (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_lsum_injective _ h.injective + +@[deprecated (since := "2024-11-24")] +alias IsInternal.submodule_independent := IsInternal.submodule_iSupIndep /-- Given an internal direct sum decomposition of a module `M`, and a basis for each of the components of the direct sum, the disjoint union of these bases is a basis for `M`. -/ @@ -374,7 +377,7 @@ the two submodules are complementary. Over a `Ring R`, this is true as an iff, a `DirectSum.isInternal_submodule_iff_isCompl`. -/ theorem IsInternal.isCompl {A : ι → Submodule R M} {i j : ι} (hij : i ≠ j) (h : (Set.univ : Set ι) = {i, j}) (hi : IsInternal A) : IsCompl (A i) (A j) := - ⟨hi.submodule_independent.pairwiseDisjoint hij, + ⟨hi.submodule_iSupIndep.pairwiseDisjoint hij, codisjoint_iff.mpr <| Eq.symm <| hi.submodule_iSup_eq_top.symm.trans <| by rw [← sSup_pair, iSup, ← Set.image_univ, h, Set.image_insert_eq, Set.image_singleton]⟩ @@ -387,60 +390,76 @@ variable {ι : Type v} [dec_ι : DecidableEq ι] variable {M : Type*} [AddCommGroup M] [Module R M] /-- Note that this is not generally true for `[Semiring R]`; see -`CompleteLattice.Independent.dfinsupp_lsum_injective` for details. -/ -theorem isInternal_submodule_of_independent_of_iSup_eq_top {A : ι → Submodule R M} - (hi : CompleteLattice.Independent A) (hs : iSup A = ⊤) : IsInternal A := +`iSupIndep.dfinsupp_lsum_injective` for details. -/ +theorem isInternal_submodule_of_iSupIndep_of_iSup_eq_top {A : ι → Submodule R M} + (hi : iSupIndep A) (hs : iSup A = ⊤) : IsInternal A := ⟨hi.dfinsupp_lsum_injective, -- Note: https://github.com/leanprover-community/mathlib4/pull/8386 had to specify value of `f` (LinearMap.range_eq_top (f := DFinsupp.lsum _ _)).1 <| (Submodule.iSup_eq_range_dfinsupp_lsum _).symm.trans hs⟩ -/-- `iff` version of `DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top`, -`DirectSum.IsInternal.submodule_independent`, and `DirectSum.IsInternal.submodule_iSup_eq_top`. -/ -theorem isInternal_submodule_iff_independent_and_iSup_eq_top (A : ι → Submodule R M) : - IsInternal A ↔ CompleteLattice.Independent A ∧ iSup A = ⊤ := - ⟨fun i ↦ ⟨i.submodule_independent, i.submodule_iSup_eq_top⟩, - And.rec isInternal_submodule_of_independent_of_iSup_eq_top⟩ +@[deprecated (since := "2024-11-24")] +alias isInternal_submodule_of_independent_of_iSup_eq_top := + isInternal_submodule_of_iSupIndep_of_iSup_eq_top + +/-- `iff` version of `DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top`, +`DirectSum.IsInternal.iSupIndep`, and `DirectSum.IsInternal.submodule_iSup_eq_top`. -/ +theorem isInternal_submodule_iff_iSupIndep_and_iSup_eq_top (A : ι → Submodule R M) : + IsInternal A ↔ iSupIndep A ∧ iSup A = ⊤ := + ⟨fun i ↦ ⟨i.submodule_iSupIndep, i.submodule_iSup_eq_top⟩, + And.rec isInternal_submodule_of_iSupIndep_of_iSup_eq_top⟩ + +@[deprecated (since := "2024-11-24")] +alias isInternal_submodule_iff_independent_and_iSup_eq_top := + isInternal_submodule_iff_iSupIndep_and_iSup_eq_top /-- If a collection of submodules has just two indices, `i` and `j`, then `DirectSum.IsInternal` is equivalent to `isCompl`. -/ theorem isInternal_submodule_iff_isCompl (A : ι → Submodule R M) {i j : ι} (hij : i ≠ j) (h : (Set.univ : Set ι) = {i, j}) : IsInternal A ↔ IsCompl (A i) (A j) := by have : ∀ k, k = i ∨ k = j := fun k ↦ by simpa using Set.ext_iff.mp h k - rw [isInternal_submodule_iff_independent_and_iSup_eq_top, iSup, ← Set.image_univ, h, - Set.image_insert_eq, Set.image_singleton, sSup_pair, CompleteLattice.independent_pair hij this] + rw [isInternal_submodule_iff_iSupIndep_and_iSup_eq_top, iSup, ← Set.image_univ, h, + Set.image_insert_eq, Set.image_singleton, sSup_pair, iSupIndep_pair hij this] exact ⟨fun ⟨hd, ht⟩ ↦ ⟨hd, codisjoint_iff.mpr ht⟩, fun ⟨hd, ht⟩ ↦ ⟨hd, ht.eq_top⟩⟩ @[simp] theorem isInternal_ne_bot_iff {A : ι → Submodule R M} : IsInternal (fun i : {i // A i ≠ ⊥} ↦ A i) ↔ IsInternal A := by - simp only [isInternal_submodule_iff_independent_and_iSup_eq_top] - exact Iff.and CompleteLattice.independent_ne_bot_iff_independent <| by simp + simp [isInternal_submodule_iff_iSupIndep_and_iSup_eq_top] -lemma isInternal_biSup_submodule_of_independent {A : ι → Submodule R M} (s : Set ι) - (h : CompleteLattice.Independent <| fun i : s ↦ A i) : +lemma isInternal_biSup_submodule_of_iSupIndep {A : ι → Submodule R M} (s : Set ι) + (h : iSupIndep <| fun i : s ↦ A i) : IsInternal <| fun (i : s) ↦ (A i).comap (⨆ i ∈ s, A i).subtype := by - refine (isInternal_submodule_iff_independent_and_iSup_eq_top _).mpr ⟨?_, by simp [iSup_subtype]⟩ + refine (isInternal_submodule_iff_iSupIndep_and_iSup_eq_top _).mpr ⟨?_, by simp [iSup_subtype]⟩ let p := ⨆ i ∈ s, A i have hp : ∀ i ∈ s, A i ≤ p := fun i hi ↦ le_biSup A hi let e : Submodule R p ≃o Set.Iic p := p.mapIic suffices (e ∘ fun i : s ↦ (A i).comap p.subtype) = fun i ↦ ⟨A i, hp i i.property⟩ by - rw [← CompleteLattice.independent_map_orderIso_iff e, this] - exact CompleteLattice.independent_of_independent_coe_Iic_comp h + rw [← iSupIndep_map_orderIso_iff e, this] + exact .of_coe_Iic_comp h ext i m change m ∈ ((A i).comap p.subtype).map p.subtype ↔ _ rw [Submodule.map_comap_subtype, inf_of_le_right (hp i i.property)] +@[deprecated (since := "2024-11-24")] +alias isInternal_biSup_submodule_of_independent := isInternal_biSup_submodule_of_iSupIndep + /-! Now copy the lemmas for subgroup and submonoids. -/ -theorem IsInternal.addSubmonoid_independent {M : Type*} [AddCommMonoid M] {A : ι → AddSubmonoid M} - (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_sumAddHom_injective _ h.injective +theorem IsInternal.addSubmonoid_iSupIndep {M : Type*} [AddCommMonoid M] {A : ι → AddSubmonoid M} + (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_sumAddHom_injective _ h.injective + +@[deprecated (since := "2024-11-24")] +alias IsInternal.addSubmonoid_independent := IsInternal.addSubmonoid_iSupIndep + +theorem IsInternal.addSubgroup_iSupIndep {G : Type*} [AddCommGroup G] {A : ι → AddSubgroup G} + (h : IsInternal A) : iSupIndep A := + iSupIndep_of_dfinsupp_sumAddHom_injective' _ h.injective -theorem IsInternal.addSubgroup_independent {M : Type*} [AddCommGroup M] {A : ι → AddSubgroup M} - (h : IsInternal A) : CompleteLattice.Independent A := - CompleteLattice.independent_of_dfinsupp_sumAddHom_injective' _ h.injective +@[deprecated (since := "2024-11-24")] +alias IsInternal.addSubgroup_independent := IsInternal.addSubgroup_iSupIndep end Ring diff --git a/Mathlib/Algebra/DirectSum/Ring.lean b/Mathlib/Algebra/DirectSum/Ring.lean index aa5a84bec944c..883dc0bb086a7 100644 --- a/Mathlib/Algebra/DirectSum/Ring.lean +++ b/Mathlib/Algebra/DirectSum/Ring.lean @@ -63,9 +63,8 @@ instances for: * `A : ι → Submodule S`: `DirectSum.GSemiring.ofSubmodules`, `DirectSum.GCommSemiring.ofSubmodules`. -If `CompleteLattice.independent (Set.range A)`, these provide a gradation of `⨆ i, A i`, and the -mapping `⨁ i, A i →+ ⨆ i, A i` can be obtained as -`DirectSum.toMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. +If `sSupIndep A`, these provide a gradation of `⨆ i, A i`, and the mapping `⨁ i, A i →+ ⨆ i, A i` +can be obtained as `DirectSum.toMonoid (fun i ↦ AddSubmonoid.inclusion <| le_iSup A i)`. ## Tags diff --git a/Mathlib/Algebra/Lie/Semisimple/Basic.lean b/Mathlib/Algebra/Lie/Semisimple/Basic.lean index 12b30d31de1c2..50b8820a1addc 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Basic.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Basic.lean @@ -147,7 +147,7 @@ lemma isSimple_of_isAtom (I : LieIdeal R L) (hI : IsAtom I) : IsSimple R I where -- Finally `⁅b, y⁆ = 0`, by the independence of the atoms. · suffices ⁅b, y.val⁆ = 0 by erw [this]; simp only [zero_mem] rw [← LieSubmodule.mem_bot (R := R) (L := L), - ← (IsSemisimple.setIndependent_isAtom hI).eq_bot] + ← (IsSemisimple.sSupIndep_isAtom hI).eq_bot] exact ⟨lie_mem_right R L I b y y.2, lie_mem_left _ _ _ _ _ hb⟩ } -- Now that we know that `J` is an ideal of `L`, -- we start with the proof that `I` is a simple Lie algebra. @@ -240,7 +240,7 @@ lemma finitelyAtomistic : ∀ s : Finset (LieIdeal R L), ↑s ⊆ {I : LieIdeal constructor -- `j` brackets to `0` with `z`, since `⁅j, z⁆` is contained in `⁅J, K⁆ ≤ J ⊓ K`, -- and `J ⊓ K = ⊥` by the independence of the atoms. - · apply (setIndependent_isAtom.disjoint_sSup (hs hJs) hs'S (Finset.not_mem_erase _ _)).le_bot + · apply (sSupIndep_isAtom.disjoint_sSup (hs hJs) hs'S (Finset.not_mem_erase _ _)).le_bot apply LieSubmodule.lie_le_inf apply LieSubmodule.lie_mem_lie j.2 simpa only [K, Finset.sup_id_eq_sSup] using hz @@ -292,7 +292,7 @@ instance (priority := 100) IsSimple.instIsSemisimple [IsSimple R L] : IsSemisimple R L := by constructor · simp - · simpa using CompleteLattice.setIndependent_singleton _ + · simpa using sSupIndep_singleton _ · intro I hI₁ hI₂ apply IsSimple.non_abelian (R := R) (L := L) rw [IsSimple.isAtom_iff_eq_top] at hI₁ diff --git a/Mathlib/Algebra/Lie/Semisimple/Defs.lean b/Mathlib/Algebra/Lie/Semisimple/Defs.lean index 018f07cacb964..74db546a94643 100644 --- a/Mathlib/Algebra/Lie/Semisimple/Defs.lean +++ b/Mathlib/Algebra/Lie/Semisimple/Defs.lean @@ -72,7 +72,7 @@ class IsSemisimple : Prop where /-- In a semisimple Lie algebra, the supremum of the atoms is the whole Lie algebra. -/ sSup_atoms_eq_top : sSup {I : LieIdeal R L | IsAtom I} = ⊤ /-- In a semisimple Lie algebra, the atoms are independent. -/ - setIndependent_isAtom : CompleteLattice.SetIndependent {I : LieIdeal R L | IsAtom I} + sSupIndep_isAtom : sSupIndep {I : LieIdeal R L | IsAtom I} /-- In a semisimple Lie algebra, the atoms are non-abelian. -/ non_abelian_of_isAtom : ∀ I : LieIdeal R L, IsAtom I → ¬ IsLieAbelian I diff --git a/Mathlib/Algebra/Lie/Submodule.lean b/Mathlib/Algebra/Lie/Submodule.lean index c6fc1d659955f..1201232787862 100644 --- a/Mathlib/Algebra/Lie/Submodule.lean +++ b/Mathlib/Algebra/Lie/Submodule.lean @@ -514,9 +514,12 @@ theorem isCompl_iff_coe_toSubmodule : IsCompl N N' ↔ IsCompl (N : Submodule R M) (N' : Submodule R M) := by simp only [isCompl_iff, disjoint_iff_coe_toSubmodule, codisjoint_iff_coe_toSubmodule] -theorem independent_iff_coe_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : - CompleteLattice.Independent N ↔ CompleteLattice.Independent fun i ↦ (N i : Submodule R M) := by - simp [CompleteLattice.independent_def, disjoint_iff_coe_toSubmodule] +theorem iSupIndep_iff_coe_toSubmodule {ι : Type*} {N : ι → LieSubmodule R L M} : + iSupIndep N ↔ iSupIndep fun i ↦ (N i : Submodule R M) := by + simp [iSupIndep_def, disjoint_iff_coe_toSubmodule] + +@[deprecated (since := "2024-11-24")] +alias independent_iff_coe_toSubmodule := iSupIndep_iff_coe_toSubmodule theorem iSup_eq_top_iff_coe_toSubmodule {ι : Sort*} {N : ι → LieSubmodule R L M} : ⨆ i, N i = ⊤ ↔ ⨆ i, (N i : Submodule R M) = ⊤ := by diff --git a/Mathlib/Algebra/Lie/TraceForm.lean b/Mathlib/Algebra/Lie/TraceForm.lean index 2d1253e2b18c9..55ee3be037f29 100644 --- a/Mathlib/Algebra/Lie/TraceForm.lean +++ b/Mathlib/Algebra/Lie/TraceForm.lean @@ -227,8 +227,8 @@ lemma traceForm_eq_sum_genWeightSpaceOf convert finite_genWeightSpaceOf_ne_bot R L M z exact LieSubmodule.coeSubmodule_eq_bot_iff (genWeightSpaceOf M _ _) classical - have h := LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpaceOf R L M z - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top h <| by + have h := LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpaceOf R L M z + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top h <| by simp [← LieSubmodule.iSup_coe_toSubmodule] simp only [LinearMap.coeFn_sum, Finset.sum_apply, traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict' hds hfin hxy] @@ -407,8 +407,8 @@ lemma traceForm_eq_sum_finrank_nsmul_mul (x y : L) : (genWeightSpace M χ) (genWeightSpace M χ) := fun χ m hm ↦ LieSubmodule.lie_mem _ <| LieSubmodule.lie_mem _ hm classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpace' K L M) + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace' K L M) (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top' K L M) simp_rw [traceForm_apply_apply, LinearMap.trace_eq_sum_trace_restrict hds hxy, ← traceForm_genWeightSpace_eq K L M _ x y] diff --git a/Mathlib/Algebra/Lie/Weights/Basic.lean b/Mathlib/Algebra/Lie/Weights/Basic.lean index ce2e34ec7f66e..8ddf44e6372c7 100644 --- a/Mathlib/Algebra/Lie/Weights/Basic.lean +++ b/Mathlib/Algebra/Lie/Weights/Basic.lean @@ -33,7 +33,7 @@ Basic definitions and properties of the above ideas are provided in this file. * `LieModule.iSup_ucs_eq_genWeightSpace_zero` * `LieModule.iInf_lowerCentralSeries_eq_posFittingComp` * `LieModule.isCompl_genWeightSpace_zero_posFittingComp` - * `LieModule.independent_genWeightSpace` + * `LieModule.iSupIndep_genWeightSpace` * `LieModule.iSup_genWeightSpace_eq_top` ## References @@ -666,32 +666,39 @@ lemma injOn_genWeightSpace [NoZeroSMulDivisors R M] : /-- Lie module weight spaces are independent. -See also `LieModule.independent_genWeightSpace'`. -/ -lemma independent_genWeightSpace [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun χ : L → R ↦ genWeightSpace M χ := by - simp only [LieSubmodule.independent_iff_coe_toSubmodule, genWeightSpace, +See also `LieModule.iSupIndep_genWeightSpace'`. -/ +lemma iSupIndep_genWeightSpace [NoZeroSMulDivisors R M] : + iSupIndep fun χ : L → R ↦ genWeightSpace M χ := by + simp only [LieSubmodule.iSupIndep_iff_coe_toSubmodule, genWeightSpace, LieSubmodule.iInf_coe_toSubmodule] exact Module.End.independent_iInf_maxGenEigenspace_of_forall_mapsTo (toEnd R L M) (fun x y φ z ↦ (genWeightSpaceOf M φ y).lie_mem) -lemma independent_genWeightSpace' [NoZeroSMulDivisors R M] : - CompleteLattice.Independent fun χ : Weight R L M ↦ genWeightSpace M χ := - (independent_genWeightSpace R L M).comp <| +@[deprecated (since := "2024-11-24")] alias independent_genWeightSpace := iSupIndep_genWeightSpace + +lemma iSupIndep_genWeightSpace' [NoZeroSMulDivisors R M] : + iSupIndep fun χ : Weight R L M ↦ genWeightSpace M χ := + (iSupIndep_genWeightSpace R L M).comp <| Subtype.val_injective.comp (Weight.equivSetOf R L M).injective -lemma independent_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : - CompleteLattice.Independent fun (χ : R) ↦ genWeightSpaceOf M χ x := by - rw [LieSubmodule.independent_iff_coe_toSubmodule] +@[deprecated (since := "2024-11-24")] alias independent_genWeightSpace' := iSupIndep_genWeightSpace' + +lemma iSupIndep_genWeightSpaceOf [NoZeroSMulDivisors R M] (x : L) : + iSupIndep fun (χ : R) ↦ genWeightSpaceOf M χ x := by + rw [LieSubmodule.iSupIndep_iff_coe_toSubmodule] dsimp [genWeightSpaceOf] exact (toEnd R L M x).independent_genEigenspace _ +@[deprecated (since := "2024-11-24")] +alias independent_genWeightSpaceOf := iSupIndep_genWeightSpaceOf + lemma finite_genWeightSpaceOf_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] (x : L) : {χ : R | genWeightSpaceOf M χ x ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpaceOf R L M x) + WellFoundedGT.finite_ne_bot_of_iSupIndep (iSupIndep_genWeightSpaceOf R L M x) lemma finite_genWeightSpace_ne_bot [NoZeroSMulDivisors R M] [IsNoetherian R M] : {χ : L → R | genWeightSpace M χ ≠ ⊥}.Finite := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent (independent_genWeightSpace R L M) + WellFoundedGT.finite_ne_bot_of_iSupIndep (iSupIndep_genWeightSpace R L M) instance Weight.instFinite [NoZeroSMulDivisors R M] [IsNoetherian R M] : Finite (Weight R L M) := by diff --git a/Mathlib/Algebra/Lie/Weights/Chain.lean b/Mathlib/Algebra/Lie/Weights/Chain.lean index 4107e1e4cd27f..1acfd102224c7 100644 --- a/Mathlib/Algebra/Lie/Weights/Chain.lean +++ b/Mathlib/Algebra/Lie/Weights/Chain.lean @@ -204,9 +204,9 @@ lemma exists_forall_mem_corootSpace_smul_add_eq_zero exact finrank_pos refine ⟨a, b, Int.ofNat_pos.mpr hb, fun x hx ↦ ?_⟩ let N : ℤ → Submodule R M := fun k ↦ genWeightSpace M (k • α + χ) - have h₁ : CompleteLattice.Independent fun (i : Finset.Ioo p q) ↦ N i := by - rw [← LieSubmodule.independent_iff_coe_toSubmodule] - refine (independent_genWeightSpace R H M).comp fun i j hij ↦ ?_ + have h₁ : iSupIndep fun (i : Finset.Ioo p q) ↦ N i := by + rw [← LieSubmodule.iSupIndep_iff_coe_toSubmodule] + refine (iSupIndep_genWeightSpace R H M).comp fun i j hij ↦ ?_ exact SetCoe.ext <| smul_left_injective ℤ hα <| by rwa [add_left_inj] at hij have h₂ : ∀ i, MapsTo (toEnd R H M x) ↑(N i) ↑(N i) := fun _ _ ↦ LieSubmodule.lie_mem _ have h₃ : genWeightSpaceChain M α χ p q = ⨆ i ∈ Finset.Ioo p q, N i := by diff --git a/Mathlib/Algebra/Lie/Weights/Killing.lean b/Mathlib/Algebra/Lie/Weights/Killing.lean index 3fe00e4149748..7c2b6ece5a6b5 100644 --- a/Mathlib/Algebra/Lie/Weights/Killing.lean +++ b/Mathlib/Algebra/Lie/Weights/Killing.lean @@ -112,8 +112,8 @@ lemma killingForm_apply_eq_zero_of_mem_rootSpace_of_add_ne_zero {α β : H → K (mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L α (β + γ) hx).comp <| mapsTo_toEnd_genWeightSpace_add_of_mem_rootSpace K L H L β γ hy classical - have hds := DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (LieSubmodule.independent_iff_coe_toSubmodule.mp <| independent_genWeightSpace K H L) + have hds := DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (LieSubmodule.iSupIndep_iff_coe_toSubmodule.mp <| iSupIndep_genWeightSpace K H L) (LieSubmodule.iSup_eq_top_iff_coe_toSubmodule.mp <| iSup_genWeightSpace_eq_top K H L) exact LinearMap.trace_eq_zero_of_mapsTo_ne hds σ hσ hf diff --git a/Mathlib/Algebra/Module/Torsion.lean b/Mathlib/Algebra/Module/Torsion.lean index 3121a844257b9..08b33bcef771a 100644 --- a/Mathlib/Algebra/Module/Torsion.lean +++ b/Mathlib/Algebra/Module/Torsion.lean @@ -96,13 +96,13 @@ theorem torsionOf_eq_bot_iff_of_noZeroSMulDivisors [Nontrivial R] [NoZeroSMulDiv · rw [mem_torsionOf_iff, smul_eq_zero] at hr tauto -/-- See also `CompleteLattice.Independent.linearIndependent` which provides the same conclusion +/-- See also `iSupIndep.linearIndependent` which provides the same conclusion but requires the stronger hypothesis `NoZeroSMulDivisors R M`. -/ -theorem CompleteLattice.Independent.linear_independent' {ι R M : Type*} {v : ι → M} [Ring R] - [AddCommGroup M] [Module R M] (hv : CompleteLattice.Independent fun i => R ∙ v i) +theorem iSupIndep.linearIndependent' {ι R M : Type*} {v : ι → M} [Ring R] + [AddCommGroup M] [Module R M] (hv : iSupIndep fun i => R ∙ v i) (h_ne_zero : ∀ i, Ideal.torsionOf R M (v i) = ⊥) : LinearIndependent R v := by refine linearIndependent_iff_not_smul_mem_span.mpr fun i r hi => ?_ - replace hv := CompleteLattice.independent_def.mp hv i + replace hv := iSupIndep_def.mp hv i simp only [iSup_subtype', ← Submodule.span_range_eq_iSup (ι := Subtype _), disjoint_iff] at hv have : r • v i ∈ (⊥ : Submodule R M) := by rw [← hv, Submodule.mem_inf] @@ -113,6 +113,9 @@ theorem CompleteLattice.Independent.linear_independent' {ι R M : Type*} {v : ι rw [← Submodule.mem_bot R, ← h_ne_zero i] simpa using this +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.linear_independent' := iSupIndep.linearIndependent' + end TorsionOf section @@ -443,8 +446,8 @@ theorem torsionBySet_isInternal {p : ι → Ideal R} (hp : (S : Set ι).Pairwise fun i j => p i ⊔ p j = ⊤) (hM : Module.IsTorsionBySet R M (⨅ i ∈ S, p i : Ideal R)) : DirectSum.IsInternal fun i : S => torsionBySet R M <| p i := - DirectSum.isInternal_submodule_of_independent_of_iSup_eq_top - (CompleteLattice.independent_iff_supIndep.mpr <| supIndep_torsionBySet_ideal hp) + DirectSum.isInternal_submodule_of_iSupIndep_of_iSup_eq_top + (iSupIndep_iff_supIndep.mpr <| supIndep_torsionBySet_ideal hp) (by apply (iSup_subtype'' ↑S fun i => torsionBySet R M <| p i).trans -- Porting note: times out if we change apply below to <| diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 4acfe80398b10..2f715eb17ed19 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -2197,9 +2197,9 @@ elements each from a different subspace in the family is linearly independent. I pairwise intersections of elements of the family are 0. -/ theorem OrthogonalFamily.independent {V : ι → Submodule 𝕜 E} (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) : - CompleteLattice.Independent V := by + iSupIndep V := by classical - apply CompleteLattice.independent_of_dfinsupp_lsum_injective + apply iSupIndep_of_dfinsupp_lsum_injective refine LinearMap.ker_eq_bot.mp ?_ rw [Submodule.eq_bot_iff] intro v hv diff --git a/Mathlib/Analysis/InnerProductSpace/Projection.lean b/Mathlib/Analysis/InnerProductSpace/Projection.lean index c2ea3c5347107..6be13bb9ce5f1 100644 --- a/Mathlib/Analysis/InnerProductSpace/Projection.lean +++ b/Mathlib/Analysis/InnerProductSpace/Projection.lean @@ -1218,7 +1218,7 @@ theorem OrthogonalFamily.isInternal_iff_of_isComplete [DecidableEq ι] {V : ι (hV : OrthogonalFamily 𝕜 (fun i => V i) fun i => (V i).subtypeₗᵢ) (hc : IsComplete (↑(iSup V) : Set E)) : DirectSum.IsInternal V ↔ (iSup V)ᗮ = ⊥ := by haveI : CompleteSpace (↥(iSup V)) := hc.completeSpace_coe - simp only [DirectSum.isInternal_submodule_iff_independent_and_iSup_eq_top, hV.independent, + simp only [DirectSum.isInternal_submodule_iff_iSupIndep_and_iSup_eq_top, hV.independent, true_and, Submodule.orthogonal_eq_bot_iff] /-- An orthogonal family of subspaces of `E` satisfies `DirectSum.IsInternal` (that is, diff --git a/Mathlib/GroupTheory/NoncommPiCoprod.lean b/Mathlib/GroupTheory/NoncommPiCoprod.lean index 3e22ae8416ac7..e5b14b7d09118 100644 --- a/Mathlib/GroupTheory/NoncommPiCoprod.lean +++ b/Mathlib/GroupTheory/NoncommPiCoprod.lean @@ -30,7 +30,7 @@ images of different morphisms commute, we obtain a canonical morphism * `MonoidHom.noncommPiCoprod_range`: The range of `MonoidHom.noncommPiCoprod` is `⨆ (i : ι), (ϕ i).range` * `Subgroup.noncommPiCoprod_range`: The range of `Subgroup.noncommPiCoprod` is `⨆ (i : ι), H i`. -* `MonoidHom.injective_noncommPiCoprod_of_independent`: in the case of groups, `pi_hom.hom` is +* `MonoidHom.injective_noncommPiCoprod_of_iSupIndep`: in the case of groups, `pi_hom.hom` is injective if the `ϕ` are injective and the ranges of the `ϕ` are independent. * `MonoidHom.independent_range_of_coprime_order`: If the `N i` have coprime orders, then the ranges of the `ϕ` are independent. @@ -48,8 +48,8 @@ variable {G : Type*} [Group G] generalizes (one direction of) `Subgroup.disjoint_iff_mul_eq_one`. -/ @[to_additive "`Finset.noncommSum` is “injective” in `f` if `f` maps into independent subgroups. This generalizes (one direction of) `AddSubgroup.disjoint_iff_add_eq_zero`. "] -theorem eq_one_of_noncommProd_eq_one_of_independent {ι : Type*} (s : Finset ι) (f : ι → G) (comm) - (K : ι → Subgroup G) (hind : CompleteLattice.Independent K) (hmem : ∀ x ∈ s, f x ∈ K x) +theorem eq_one_of_noncommProd_eq_one_of_iSupIndep {ι : Type*} (s : Finset ι) (f : ι → G) (comm) + (K : ι → Subgroup G) (hind : iSupIndep K) (hmem : ∀ x ∈ s, f x ∈ K x) (heq1 : s.noncommProd f comm = 1) : ∀ i ∈ s, f i = 1 := by classical revert heq1 @@ -73,6 +73,9 @@ theorem eq_one_of_noncommProd_eq_one_of_independent {ι : Type*} (s : Finset ι) · exact heq1i · refine ih hcomm hmem.2 heq1S _ h +@[deprecated (since := "2024-11-24")] +alias eq_one_of_noncommProd_eq_one_of_independent := eq_one_of_noncommProd_eq_one_of_iSupIndep + end Subgroup section FamilyOfMonoids @@ -200,27 +203,30 @@ theorem noncommPiCoprod_range [Fintype ι] exact ⟨Pi.mulSingle i y, noncommPiCoprod_mulSingle _ _ _⟩ @[to_additive] -theorem injective_noncommPiCoprod_of_independent [Fintype ι] +theorem injective_noncommPiCoprod_of_iSupIndep [Fintype ι] {hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)} - (hind : CompleteLattice.Independent fun i => (ϕ i).range) + (hind : iSupIndep fun i => (ϕ i).range) (hinj : ∀ i, Function.Injective (ϕ i)) : Function.Injective (noncommPiCoprod ϕ hcomm) := by classical apply (MonoidHom.ker_eq_bot_iff _).mp rw [eq_bot_iff] intro f heq1 have : ∀ i, i ∈ Finset.univ → ϕ i (f i) = 1 := - Subgroup.eq_one_of_noncommProd_eq_one_of_independent _ _ (fun _ _ _ _ h => hcomm h _ _) + Subgroup.eq_one_of_noncommProd_eq_one_of_iSupIndep _ _ (fun _ _ _ _ h => hcomm h _ _) _ hind (by simp) heq1 ext i apply hinj simp [this i (Finset.mem_univ i)] +@[deprecated (since := "2024-11-24")] +alias injective_noncommPiCoprod_of_independent := injective_noncommPiCoprod_of_iSupIndep + @[to_additive] theorem independent_range_of_coprime_order (hcomm : Pairwise fun i j : ι => ∀ (x : H i) (y : H j), Commute (ϕ i x) (ϕ j y)) [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : - CompleteLattice.Independent fun i => (ϕ i).range := by + iSupIndep fun i => (ϕ i).range := by cases nonempty_fintype ι letI := Classical.decEq ι rintro i @@ -279,7 +285,7 @@ theorem independent_of_coprime_order (hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y) [Finite ι] [∀ i, Fintype (H i)] (hcoprime : Pairwise fun i j => Nat.Coprime (Fintype.card (H i)) (Fintype.card (H j))) : - CompleteLattice.Independent H := by + iSupIndep H := by simpa using MonoidHom.independent_range_of_coprime_order (fun i => (H i).subtype) (commute_subtype_of_commute hcomm) hcoprime @@ -306,11 +312,11 @@ theorem noncommPiCoprod_range simp [noncommPiCoprod, MonoidHom.noncommPiCoprod_range] @[to_additive] -theorem injective_noncommPiCoprod_of_independent +theorem injective_noncommPiCoprod_of_iSupIndep {hcomm : Pairwise fun i j : ι => ∀ x y : G, x ∈ H i → y ∈ H j → Commute x y} - (hind : CompleteLattice.Independent H) : + (hind : iSupIndep H) : Function.Injective (noncommPiCoprod hcomm) := by - apply MonoidHom.injective_noncommPiCoprod_of_independent + apply MonoidHom.injective_noncommPiCoprod_of_iSupIndep · simpa using hind · intro i exact Subtype.coe_injective diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index 5521719559622..fdfa524a631b9 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -806,7 +806,7 @@ noncomputable def directProductOfNormal [Finite G] apply MulEquiv.ofBijective (Subgroup.noncommPiCoprod hcomm) apply (Fintype.bijective_iff_injective_and_card _).mpr constructor - · apply Subgroup.injective_noncommPiCoprod_of_independent + · apply Subgroup.injective_noncommPiCoprod_of_iSupIndep apply independent_of_coprime_order hcomm rintro ⟨p₁, hp₁⟩ ⟨p₂, hp₂⟩ hne haveI hp₁' := Fact.mk (Nat.prime_of_mem_primeFactors hp₁) diff --git a/Mathlib/LinearAlgebra/DFinsupp.lean b/Mathlib/LinearAlgebra/DFinsupp.lean index 406aa08a4c032..f411c6bc0c1b5 100644 --- a/Mathlib/LinearAlgebra/DFinsupp.lean +++ b/Mathlib/LinearAlgebra/DFinsupp.lean @@ -390,8 +390,6 @@ theorem mem_iSup_finset_iff_exists_sum {s : Finset ι} (p : ι → Submodule R N end Submodule -namespace CompleteLattice - open DFinsupp section Semiring @@ -401,23 +399,26 @@ variable [DecidableEq ι] [Semiring R] [AddCommMonoid N] [Module R N] /-- Independence of a family of submodules can be expressed as a quantifier over `DFinsupp`s. This is an intermediate result used to prove -`CompleteLattice.independent_of_dfinsupp_lsum_injective` and -`CompleteLattice.Independent.dfinsupp_lsum_injective`. -/ -theorem independent_iff_forall_dfinsupp (p : ι → Submodule R N) : - Independent p ↔ +`iSupIndep_of_dfinsupp_lsum_injective` and +`iSupIndep.dfinsupp_lsum_injective`. -/ +theorem iSupIndep_iff_forall_dfinsupp (p : ι → Submodule R N) : + iSupIndep p ↔ ∀ (i) (x : p i) (v : Π₀ i : ι, ↥(p i)), lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (erase i v) = x → x = 0 := by - simp_rw [CompleteLattice.independent_def, Submodule.disjoint_def, + simp_rw [iSupIndep_def, Submodule.disjoint_def, Submodule.mem_biSup_iff_exists_dfinsupp, exists_imp, filter_ne_eq_erase] refine forall_congr' fun i => Subtype.forall'.trans ?_ simp_rw [Submodule.coe_eq_zero] +@[deprecated (since := "2024-11-24")] +alias independent_iff_forall_dfinsupp := iSupIndep_iff_forall_dfinsupp + /- If `DFinsupp.lsum` applied with `Submodule.subtype` is injective then the submodules are -independent. -/ -theorem independent_of_dfinsupp_lsum_injective (p : ι → Submodule R N) +iSupIndep. -/ +theorem iSupIndep_of_dfinsupp_lsum_injective (p : ι → Submodule R N) (h : Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype)) : - Independent p := by - rw [independent_iff_forall_dfinsupp] + iSupIndep p := by + rw [iSupIndep_iff_forall_dfinsupp] intro i x v hv replace hv : lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (erase i v) = lsum ℕ (M := fun i ↦ ↥(p i)) (fun i => (p i).subtype) (single i x) := by @@ -425,12 +426,18 @@ theorem independent_of_dfinsupp_lsum_injective (p : ι → Submodule R N) have := DFunLike.ext_iff.mp (h hv) i simpa [eq_comm] using this +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_lsum_injective := iSupIndep_of_dfinsupp_lsum_injective + /- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive submonoids are independent. -/ -theorem independent_of_dfinsupp_sumAddHom_injective (p : ι → AddSubmonoid N) - (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : Independent p := by - rw [← independent_map_orderIso_iff (AddSubmonoid.toNatSubmodule : AddSubmonoid N ≃o _)] - exact independent_of_dfinsupp_lsum_injective _ h +theorem iSupIndep_of_dfinsupp_sumAddHom_injective (p : ι → AddSubmonoid N) + (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : iSupIndep p := by + rw [← iSupIndep_map_orderIso_iff (AddSubmonoid.toNatSubmodule : AddSubmonoid N ≃o _)] + exact iSupIndep_of_dfinsupp_lsum_injective _ h + +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_sumAddHom_injective := iSupIndep_of_dfinsupp_sumAddHom_injective /-- Combining `DFinsupp.lsum` with `LinearMap.toSpanSingleton` is the same as `Finsupp.linearCombination` -/ @@ -450,24 +457,27 @@ section Ring variable [DecidableEq ι] [Ring R] [AddCommGroup N] [Module R N] -/- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive +/-- If `DFinsupp.sumAddHom` applied with `AddSubmonoid.subtype` is injective then the additive subgroups are independent. -/ -theorem independent_of_dfinsupp_sumAddHom_injective' (p : ι → AddSubgroup N) - (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : Independent p := by - rw [← independent_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] - exact independent_of_dfinsupp_lsum_injective _ h +theorem iSupIndep_of_dfinsupp_sumAddHom_injective' (p : ι → AddSubgroup N) + (h : Function.Injective (sumAddHom fun i => (p i).subtype)) : iSupIndep p := by + rw [← iSupIndep_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] + exact iSupIndep_of_dfinsupp_lsum_injective _ h + +@[deprecated (since := "2024-11-24")] +alias independent_of_dfinsupp_sumAddHom_injective' := iSupIndep_of_dfinsupp_sumAddHom_injective' /-- The canonical map out of a direct sum of a family of submodules is injective when the submodules -are `CompleteLattice.Independent`. +are `iSupIndep`. Note that this is not generally true for `[Semiring R]`, for instance when `A` is the `ℕ`-submodules of the positive and negative integers. See `Counterexamples/DirectSumIsInternal.lean` for a proof of this fact. -/ -theorem Independent.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : Independent p) : +theorem iSupIndep.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : iSupIndep p) : Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := by -- simplify everything down to binders over equalities in `N` - rw [independent_iff_forall_dfinsupp] at h + rw [iSupIndep_iff_forall_dfinsupp] at h suffices LinearMap.ker (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) = ⊥ by -- Lean can't find this without our help letI thisI : AddCommGroup (Π₀ i, p i) := inferInstance @@ -482,34 +492,46 @@ theorem Independent.dfinsupp_lsum_injective {p : ι → Submodule R N} (h : Inde rwa [← erase_add_single i m, LinearMap.map_add, lsum_single, Submodule.subtype_apply, add_eq_zero_iff_eq_neg, ← Submodule.coe_neg] at hm +@[deprecated (since := "2024-11-24")] +alias Independent.dfinsupp_lsum_injective := iSupIndep.dfinsupp_lsum_injective + /-- The canonical map out of a direct sum of a family of additive subgroups is injective when the -additive subgroups are `CompleteLattice.Independent`. -/ -theorem Independent.dfinsupp_sumAddHom_injective {p : ι → AddSubgroup N} (h : Independent p) : +additive subgroups are `iSupIndep`. -/ +theorem iSupIndep.dfinsupp_sumAddHom_injective {p : ι → AddSubgroup N} (h : iSupIndep p) : Function.Injective (sumAddHom fun i => (p i).subtype) := by - rw [← independent_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] at h + rw [← iSupIndep_map_orderIso_iff (AddSubgroup.toIntSubmodule : AddSubgroup N ≃o _)] at h exact h.dfinsupp_lsum_injective +@[deprecated (since := "2024-11-24")] +alias Independent.dfinsupp_sumAddHom_injective := iSupIndep.dfinsupp_sumAddHom_injective + /-- A family of submodules over an additive group are independent if and only iff `DFinsupp.lsum` applied with `Submodule.subtype` is injective. Note that this is not generally true for `[Semiring R]`; see -`CompleteLattice.Independent.dfinsupp_lsum_injective` for details. -/ -theorem independent_iff_dfinsupp_lsum_injective (p : ι → Submodule R N) : - Independent p ↔ Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := - ⟨Independent.dfinsupp_lsum_injective, independent_of_dfinsupp_lsum_injective p⟩ +`iSupIndep.dfinsupp_lsum_injective` for details. -/ +theorem iSupIndep_iff_dfinsupp_lsum_injective (p : ι → Submodule R N) : + iSupIndep p ↔ Function.Injective (lsum ℕ (M := fun i ↦ ↥(p i)) fun i => (p i).subtype) := + ⟨iSupIndep.dfinsupp_lsum_injective, iSupIndep_of_dfinsupp_lsum_injective p⟩ + +@[deprecated (since := "2024-11-24")] +alias independent_iff_dfinsupp_lsum_injective := iSupIndep_iff_dfinsupp_lsum_injective /-- A family of additive subgroups over an additive group are independent if and only if `DFinsupp.sumAddHom` applied with `AddSubgroup.subtype` is injective. -/ -theorem independent_iff_dfinsupp_sumAddHom_injective (p : ι → AddSubgroup N) : - Independent p ↔ Function.Injective (sumAddHom fun i => (p i).subtype) := - ⟨Independent.dfinsupp_sumAddHom_injective, independent_of_dfinsupp_sumAddHom_injective' p⟩ +theorem iSupIndep_iff_dfinsupp_sumAddHom_injective (p : ι → AddSubgroup N) : + iSupIndep p ↔ Function.Injective (sumAddHom fun i => (p i).subtype) := + ⟨iSupIndep.dfinsupp_sumAddHom_injective, iSupIndep_of_dfinsupp_sumAddHom_injective' p⟩ + +@[deprecated (since := "2024-11-24")] +alias independent_iff_dfinsupp_sumAddHom_injective := iSupIndep_iff_dfinsupp_sumAddHom_injective -/-- If a family of submodules is `Independent`, then a choice of nonzero vector from each submodule +/-- If a family of submodules is independent, then a choice of nonzero vector from each submodule forms a linearly independent family. -See also `CompleteLattice.Independent.linearIndependent'`. -/ -theorem Independent.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → Submodule R N) - (hp : Independent p) {v : ι → N} (hv : ∀ i, v i ∈ p i) (hv' : ∀ i, v i ≠ 0) : +See also `iSupIndep.linearIndependent'`. -/ +theorem iSupIndep.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → Submodule R N) + (hp : iSupIndep p) {v : ι → N} (hv : ∀ i, v i ∈ p i) (hv' : ∀ i, v i ≠ 0) : LinearIndependent R v := by let _ := Classical.decEq ι let _ := Classical.decEq R @@ -527,15 +549,19 @@ theorem Independent.linearIndependent [NoZeroSMulDivisors R N] {ι} (p : ι → simp only [coe_zero, Pi.zero_apply, ZeroMemClass.coe_zero, smul_eq_zero, ha] at this simpa -theorem independent_iff_linearIndependent_of_ne_zero [NoZeroSMulDivisors R N] {ι} {v : ι → N} - (h_ne_zero : ∀ i, v i ≠ 0) : (Independent fun i => R ∙ v i) ↔ LinearIndependent R v := +@[deprecated (since := "2024-11-24")] +alias Independent.linearIndependent := iSupIndep.linearIndependent + +theorem iSupIndep_iff_linearIndependent_of_ne_zero [NoZeroSMulDivisors R N] {ι} {v : ι → N} + (h_ne_zero : ∀ i, v i ≠ 0) : (iSupIndep fun i => R ∙ v i) ↔ LinearIndependent R v := let _ := Classical.decEq ι ⟨fun hv => hv.linearIndependent _ (fun i => Submodule.mem_span_singleton_self <| v i) h_ne_zero, - fun hv => hv.independent_span_singleton⟩ + fun hv => hv.iSupIndep_span_singleton⟩ -end Ring +@[deprecated (since := "2024-11-24")] +alias independent_iff_linearIndependent_of_ne_zero := iSupIndep_iff_linearIndependent_of_ne_zero -end CompleteLattice +end Ring namespace LinearMap diff --git a/Mathlib/LinearAlgebra/Dimension/Finite.lean b/Mathlib/LinearAlgebra/Dimension/Finite.lean index 49fba6493d2d1..6c000a000aaf9 100644 --- a/Mathlib/LinearAlgebra/Dimension/Finite.lean +++ b/Mathlib/LinearAlgebra/Dimension/Finite.lean @@ -249,8 +249,8 @@ theorem Module.Finite.not_linearIndependent_of_infinite {ι : Type*} [Infinite section variable [NoZeroSMulDivisors R M] -theorem CompleteLattice.Independent.subtype_ne_bot_le_rank [Nontrivial R] - {V : ι → Submodule R M} (hV : CompleteLattice.Independent V) : +theorem iSupIndep.subtype_ne_bot_le_rank [Nontrivial R] + {V : ι → Submodule R M} (hV : iSupIndep V) : Cardinal.lift.{v} #{ i : ι // V i ≠ ⊥ } ≤ Cardinal.lift.{w} (Module.rank R M) := by set I := { i : ι // V i ≠ ⊥ } have hI : ∀ i : I, ∃ v ∈ V i, v ≠ (0 : M) := by @@ -261,10 +261,13 @@ theorem CompleteLattice.Independent.subtype_ne_bot_le_rank [Nontrivial R] have : LinearIndependent R v := (hV.comp Subtype.coe_injective).linearIndependent _ hvV hv exact this.cardinal_lift_le_rank +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.subtype_ne_bot_le_rank := iSupIndep.subtype_ne_bot_le_rank + variable [Module.Finite R M] [StrongRankCondition R] -theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) : +theorem iSupIndep.subtype_ne_bot_le_finrank_aux + {p : ι → Submodule R M} (hp : iSupIndep p) : #{ i // p i ≠ ⊥ } ≤ (finrank R M : Cardinal.{w}) := by suffices Cardinal.lift.{v} #{ i // p i ≠ ⊥ } ≤ Cardinal.lift.{v} (finrank R M : Cardinal.{w}) by rwa [Cardinal.lift_le] at this @@ -276,8 +279,8 @@ theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank_aux /-- If `p` is an independent family of submodules of a `R`-finite module `M`, then the number of nontrivial subspaces in the family `p` is finite. -/ -noncomputable def CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) : +noncomputable def iSupIndep.fintypeNeBotOfFiniteDimensional + {p : ι → Submodule R M} (hp : iSupIndep p) : Fintype { i : ι // p i ≠ ⊥ } := by suffices #{ i // p i ≠ ⊥ } < (ℵ₀ : Cardinal.{w}) by rw [Cardinal.lt_aleph0_iff_fintype] at this @@ -289,9 +292,9 @@ noncomputable def CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional number of nontrivial subspaces in the family `p` is bounded above by the dimension of `M`. Note that the `Fintype` hypothesis required here can be provided by -`CompleteLattice.Independent.fintypeNeBotOfFiniteDimensional`. -/ -theorem CompleteLattice.Independent.subtype_ne_bot_le_finrank - {p : ι → Submodule R M} (hp : CompleteLattice.Independent p) [Fintype { i // p i ≠ ⊥ }] : +`iSupIndep.fintypeNeBotOfFiniteDimensional`. -/ +theorem iSupIndep.subtype_ne_bot_le_finrank + {p : ι → Submodule R M} (hp : iSupIndep p) [Fintype { i // p i ≠ ⊥ }] : Fintype.card { i // p i ≠ ⊥ } ≤ finrank R M := by simpa using hp.subtype_ne_bot_le_finrank_aux end diff --git a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean index 05c1324d366e9..3003ced8d18b7 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Basic.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Basic.lean @@ -664,11 +664,11 @@ lemma injOn_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : apply injOn_maxGenEigenspace theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : ℕ∞) : - CompleteLattice.Independent (f.genEigenspace · k) := by + iSupIndep (f.genEigenspace · k) := by classical suffices ∀ μ₁ (s : Finset R), μ₁ ∉ s → Disjoint (f.genEigenspace μ₁ k) (s.sup fun μ ↦ f.genEigenspace μ k) by - simp_rw [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_genEigenspace f k), + simp_rw [iSupIndep_iff_supIndep_of_injOn (injOn_genEigenspace f k), Finset.supIndep_iff_disjoint_erase] exact fun s μ _ ↦ this _ _ (s.not_mem_erase μ) intro μ₁ s @@ -703,27 +703,29 @@ theorem independent_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) (k : rwa [ih.eq_bot, Submodule.mem_bot] at hyz theorem independent_maxGenEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent f.maxGenEigenspace := by + iSupIndep f.maxGenEigenspace := by apply independent_genEigenspace @[deprecated independent_genEigenspace (since := "2024-10-23")] theorem independent_iSup_genEigenspace [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent (fun μ ↦ ⨆ k : ℕ, f.genEigenspace μ k) := by + iSupIndep (fun μ ↦ ⨆ k : ℕ, f.genEigenspace μ k) := by simp_rw [iSup_genEigenspace_eq] apply independent_maxGenEigenspace /-- The eigenspaces of a linear operator form an independent family of subspaces of `M`. That is, any eigenspace has trivial intersection with the span of all the other eigenspaces. -/ -theorem eigenspaces_independent [NoZeroSMulDivisors R M] (f : End R M) : - CompleteLattice.Independent f.eigenspace := +theorem eigenspaces_iSupIndep [NoZeroSMulDivisors R M] (f : End R M) : + iSupIndep f.eigenspace := (f.independent_genEigenspace 1).mono fun _ ↦ le_rfl +@[deprecated (since := "2024-11-24")] alias eigenspaces_independent := eigenspaces_iSupIndep + /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly independent. -/ theorem eigenvectors_linearIndependent' {ι : Type*} [NoZeroSMulDivisors R M] (f : End R M) (μ : ι → R) (hμ : Function.Injective μ) (v : ι → M) (h_eigenvec : ∀ i, f.HasEigenvector (μ i) (v i)) : LinearIndependent R v := - f.eigenspaces_independent.comp hμ |>.linearIndependent _ + f.eigenspaces_iSupIndep.comp hμ |>.linearIndependent _ (fun i ↦ h_eigenvec i |>.left) (fun i ↦ h_eigenvec i |>.right) /-- Eigenvectors corresponding to distinct eigenvalues of a linear operator are linearly diff --git a/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean b/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean index 7ae63cc2e199b..0b327ae0fd03d 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Matrix.lean @@ -55,7 +55,7 @@ lemma hasEigenvalue_toLin_diagonal_iff (d : n → R) {μ : R} [NoZeroSMulDivisor rw [mem_eigenspace_iff] exact (hasEigenvector_toLin_diagonal d i b).apply_eq_smul have hμ_not_mem : μ ∉ Set.range d := by simpa using fun i ↦ (hμ i) - have := eigenspaces_independent (toLin b b (diagonal d)) |>.disjoint_biSup hμ_not_mem + have := eigenspaces_iSupIndep (toLin b b (diagonal d)) |>.disjoint_biSup hμ_not_mem rw [h_iSup, disjoint_top] at this exact h_eig this · rintro ⟨i, rfl⟩ diff --git a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean index 9f089022d24ad..ddb8f9f9ea5ac 100644 --- a/Mathlib/LinearAlgebra/Eigenspace/Pi.lean +++ b/Mathlib/LinearAlgebra/Eigenspace/Pi.lean @@ -85,7 +85,7 @@ lemma injOn_iInf_maxGenEigenspace : lemma independent_iInf_maxGenEigenspace_of_forall_mapsTo (h : ∀ i j φ, MapsTo (f i) ((f j).maxGenEigenspace φ) ((f j).maxGenEigenspace φ)) : - CompleteLattice.Independent fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i) := by + iSupIndep fun χ : ι → R ↦ ⨅ i, (f i).maxGenEigenspace (χ i) := by replace h (l : ι) (χ : ι → R) : MapsTo (f l) (⨅ i, (f i).maxGenEigenspace (χ i)) (⨅ i, (f i).maxGenEigenspace (χ i)) := by intro x hx @@ -95,7 +95,7 @@ lemma independent_iInf_maxGenEigenspace_of_forall_mapsTo suffices ∀ χ (s : Finset (ι → R)) (_ : χ ∉ s), Disjoint (⨅ i, (f i).maxGenEigenspace (χ i)) (s.sup fun (χ : ι → R) ↦ ⨅ i, (f i).maxGenEigenspace (χ i)) by - simpa only [CompleteLattice.independent_iff_supIndep_of_injOn (injOn_iInf_maxGenEigenspace f), + simpa only [iSupIndep_iff_supIndep_of_injOn (injOn_iInf_maxGenEigenspace f), Finset.supIndep_iff_disjoint_erase] using fun s χ _ ↦ this _ _ (s.not_mem_erase χ) intro χ₁ s induction s using Finset.induction_on with diff --git a/Mathlib/LinearAlgebra/LinearIndependent.lean b/Mathlib/LinearAlgebra/LinearIndependent.lean index 39d2c59e03232..eedfde103f8a0 100644 --- a/Mathlib/LinearAlgebra/LinearIndependent.lean +++ b/Mathlib/LinearAlgebra/LinearIndependent.lean @@ -913,10 +913,10 @@ theorem linearIndependent_iff_not_smul_mem_span : simp [hij] · simp [hl]⟩ -/-- See also `CompleteLattice.independent_iff_linearIndependent_of_ne_zero`. -/ -theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v) : - CompleteLattice.Independent fun i => R ∙ v i := by - refine CompleteLattice.independent_def.mp fun i => ?_ +/-- See also `iSupIndep_iff_linearIndependent_of_ne_zero`. -/ +theorem LinearIndependent.iSupIndep_span_singleton (hv : LinearIndependent R v) : + iSupIndep fun i => R ∙ v i := by + refine iSupIndep_def.mp fun i => ?_ rw [disjoint_iff_inf_le] intro m hm simp only [mem_inf, mem_span_singleton, iSup_subtype'] at hm @@ -928,6 +928,9 @@ theorem LinearIndependent.independent_span_singleton (hv : LinearIndependent R v ext simp +@[deprecated (since := "2024-11-24")] +alias LinearIndependent.independent_span_singleton := LinearIndependent.iSupIndep_span_singleton + variable (R) theorem exists_maximal_independent' (s : ι → M) : diff --git a/Mathlib/LinearAlgebra/Projectivization/Independence.lean b/Mathlib/LinearAlgebra/Projectivization/Independence.lean index e9a6bf429c7e2..403138f8a84d9 100644 --- a/Mathlib/LinearAlgebra/Projectivization/Independence.lean +++ b/Mathlib/LinearAlgebra/Projectivization/Independence.lean @@ -56,17 +56,19 @@ theorem independent_iff : Independent f ↔ LinearIndependent K (Projectivizatio /-- A family of points in projective space is independent if and only if the family of submodules which the points determine is independent in the lattice-theoretic sense. -/ -theorem independent_iff_completeLattice_independent : - Independent f ↔ CompleteLattice.Independent fun i => (f i).submodule := by +theorem independent_iff_iSupIndep : Independent f ↔ iSupIndep fun i => (f i).submodule := by refine ⟨?_, fun h => ?_⟩ · rintro ⟨f, hf, hi⟩ simp only [submodule_mk] - exact (CompleteLattice.independent_iff_linearIndependent_of_ne_zero (R := K) hf).mpr hi + exact (iSupIndep_iff_linearIndependent_of_ne_zero (R := K) hf).mpr hi · rw [independent_iff] refine h.linearIndependent (Projectivization.submodule ∘ f) (fun i => ?_) fun i => ?_ · simpa only [Function.comp_apply, submodule_eq] using Submodule.mem_span_singleton_self _ · exact rep_nonzero (f i) +@[deprecated (since := "2024-11-24")] +alias independent_iff_completeLattice_independent := independent_iff_iSupIndep + /-- A linearly dependent family of nonzero vectors gives a dependent family of points in projective space. -/ inductive Dependent : (ι → ℙ K V) → Prop diff --git a/Mathlib/Order/CompactlyGenerated/Basic.lean b/Mathlib/Order/CompactlyGenerated/Basic.lean index f3bc338a1d2ea..e802e01e3fb3e 100644 --- a/Mathlib/Order/CompactlyGenerated/Basic.lean +++ b/Mathlib/Order/CompactlyGenerated/Basic.lean @@ -273,13 +273,14 @@ alias ⟨_, IsSupClosedCompact.isSupFiniteCompact⟩ := isSupFiniteCompact_iff_i alias ⟨_, WellFoundedGT.isSupClosedCompact⟩ := isSupClosedCompact_iff_wellFoundedGT -variable {α} +end CompleteLattice + -theorem WellFoundedGT.finite_of_setIndependent [WellFoundedGT α] {s : Set α} - (hs : SetIndependent s) : s.Finite := by +theorem WellFoundedGT.finite_of_sSupIndep [WellFoundedGT α] {s : Set α} + (hs : sSupIndep s) : s.Finite := by classical refine Set.not_infinite.mp fun contra => ?_ - obtain ⟨t, ht₁, ht₂⟩ := WellFoundedGT.isSupFiniteCompact α s + obtain ⟨t, ht₁, ht₂⟩ := CompleteLattice.WellFoundedGT.isSupFiniteCompact α s replace contra : ∃ x : α, x ∈ s ∧ x ≠ ⊥ ∧ x ∉ t := by have : (s \ (insert ⊥ t : Finset α)).Infinite := contra.diff (Finset.finite_toSet _) obtain ⟨x, hx₁, hx₂⟩ := this.nonempty @@ -290,41 +291,59 @@ theorem WellFoundedGT.finite_of_setIndependent [WellFoundedGT α] {s : Set α} simpa [Disjoint, hx₂, ← t.sup_id_eq_sSup, ← ht₂] using this.eq_bot apply hx₁ rw [← hs, eq_comm, inf_eq_left] - exact le_sSup _ _ hx₀ + exact le_sSup hx₀ -theorem WellFoundedGT.finite_ne_bot_of_independent [WellFoundedGT α] - {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_of_setIndependent := WellFoundedGT.finite_of_sSupIndep + +theorem WellFoundedGT.finite_ne_bot_of_iSupIndep [WellFoundedGT α] + {ι : Type*} {t : ι → α} (ht : iSupIndep t) : Set.Finite {i | t i ≠ ⊥} := by refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn - exact WellFoundedGT.finite_of_setIndependent ht.setIndependent_range + exact WellFoundedGT.finite_of_sSupIndep ht.sSupIndep_range + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent := + WellFoundedGT.finite_ne_bot_of_iSupIndep -theorem WellFoundedGT.finite_of_independent [WellFoundedGT α] {ι : Type*} - {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := - haveI := (WellFoundedGT.finite_of_setIndependent ht.setIndependent_range).to_subtype +theorem WellFoundedGT.finite_of_iSupIndep [WellFoundedGT α] {ι : Type*} + {t : ι → α} (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := + haveI := (WellFoundedGT.finite_of_sSupIndep ht.sSupIndep_range).to_subtype Finite.of_injective_finite_range (ht.injective h_ne_bot) -theorem WellFoundedLT.finite_of_setIndependent [WellFoundedLT α] {s : Set α} - (hs : SetIndependent s) : s.Finite := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedGT.finite_of_independent := WellFoundedGT.finite_of_iSupIndep + +theorem WellFoundedLT.finite_of_sSupIndep [WellFoundedLT α] {s : Set α} + (hs : sSupIndep s) : s.Finite := by by_contra inf let e := (Infinite.diff inf <| finite_singleton ⊥).to_subtype.natEmbedding let a n := ⨆ i ≥ n, (e i).1 have sup_le n : (e n).1 ⊔ a (n + 1) ≤ a n := sup_le_iff.mpr ⟨le_iSup₂_of_le n le_rfl le_rfl, iSup₂_le fun i hi ↦ le_iSup₂_of_le i (n.le_succ.trans hi) le_rfl⟩ have lt n : a (n + 1) < a n := (Disjoint.right_lt_sup_of_left_ne_bot - ((hs (e n).2.1).mono_right <| iSup₂_le fun i hi ↦ le_sSup _ _ ?_) (e n).2.2).trans_le (sup_le n) + ((hs (e n).2.1).mono_right <| iSup₂_le fun i hi ↦ le_sSup ?_) (e n).2.2).trans_le (sup_le n) · exact (RelEmbedding.natGT a lt).not_wellFounded_of_decreasing_seq wellFounded_lt exact ⟨(e i).2.1, fun h ↦ n.lt_succ_self.not_le <| hi.trans_eq <| e.2 <| Subtype.val_injective h⟩ -theorem WellFoundedLT.finite_ne_bot_of_independent [WellFoundedLT α] - {ι : Type*} {t : ι → α} (ht : Independent t) : Set.Finite {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_of_setIndependent := WellFoundedLT.finite_of_sSupIndep + +theorem WellFoundedLT.finite_ne_bot_of_iSupIndep [WellFoundedLT α] + {ι : Type*} {t : ι → α} (ht : iSupIndep t) : Set.Finite {i | t i ≠ ⊥} := by refine Finite.of_finite_image (Finite.subset ?_ (image_subset_range t _)) ht.injOn - exact WellFoundedLT.finite_of_setIndependent ht.setIndependent_range + exact WellFoundedLT.finite_of_sSupIndep ht.sSupIndep_range + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_ne_bot_of_independent := + WellFoundedLT.finite_ne_bot_of_iSupIndep -theorem WellFoundedLT.finite_of_independent [WellFoundedLT α] {ι : Type*} - {t : ι → α} (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := - haveI := (WellFoundedLT.finite_of_setIndependent ht.setIndependent_range).to_subtype +theorem WellFoundedLT.finite_of_iSupIndep [WellFoundedLT α] {ι : Type*} + {t : ι → α} (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Finite ι := + haveI := (WellFoundedLT.finite_of_sSupIndep ht.sSupIndep_range).to_subtype Finite.of_injective_finite_range (ht.injective h_ne_bot) -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.WellFoundedLT.finite_of_independent := WellFoundedLT.finite_of_iSupIndep /-- A complete lattice is said to be compactly generated if any element is the `sSup` of compact elements. -/ @@ -413,9 +432,9 @@ theorem inf_sSup_eq_iSup_inf_sup_finset : (iSup_le fun t => iSup_le fun h => inf_le_inf_left _ ((Finset.sup_id_eq_sSup t).symm ▸ sSup_le_sSup h)) -theorem CompleteLattice.setIndependent_iff_finite {s : Set α} : - CompleteLattice.SetIndependent s ↔ - ∀ t : Finset α, ↑t ⊆ s → CompleteLattice.SetIndependent (↑t : Set α) := +theorem sSupIndep_iff_finite {s : Set α} : + sSupIndep s ↔ + ∀ t : Finset α, ↑t ⊆ s → sSupIndep (↑t : Set α) := ⟨fun hs _ ht => hs.mono ht, fun h a ha => by rw [disjoint_iff, inf_sSup_eq_iSup_inf_sup_finset, iSup_eq_bot] intro t @@ -428,10 +447,13 @@ theorem CompleteLattice.setIndependent_iff_finite {s : Set α} : · rw [Finset.coe_insert, Set.insert_subset_iff] exact ⟨ha, Set.Subset.trans ht diff_subset⟩⟩ -lemma CompleteLattice.independent_iff_supIndep_of_injOn {ι : Type*} {f : ι → α} +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_iff_finite := sSupIndep_iff_finite + +lemma iSupIndep_iff_supIndep_of_injOn {ι : Type*} {f : ι → α} (hf : InjOn f {i | f i ≠ ⊥}) : - CompleteLattice.Independent f ↔ ∀ (s : Finset ι), s.SupIndep f := by - refine ⟨fun h ↦ h.supIndep', fun h ↦ CompleteLattice.independent_def'.mpr fun i ↦ ?_⟩ + iSupIndep f ↔ ∀ (s : Finset ι), s.SupIndep f := by + refine ⟨fun h ↦ h.supIndep', fun h ↦ iSupIndep_def'.mpr fun i ↦ ?_⟩ simp_rw [disjoint_iff, inf_sSup_eq_iSup_inf_sup_finset, iSup_eq_bot, ← disjoint_iff] intro s hs classical @@ -451,11 +473,14 @@ lemma CompleteLattice.independent_iff_supIndep_of_injOn {ι : Type*} {f : ι → rw [Finset.supIndep_iff_disjoint_erase] at h exact h i (Finset.mem_insert_self i _) -theorem CompleteLattice.setIndependent_iUnion_of_directed {η : Type*} {s : η → Set α} - (hs : Directed (· ⊆ ·) s) (h : ∀ i, CompleteLattice.SetIndependent (s i)) : - CompleteLattice.SetIndependent (⋃ i, s i) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep_of_injOn := iSupIndep_iff_supIndep_of_injOn + +theorem sSupIndep_iUnion_of_directed {η : Type*} {s : η → Set α} + (hs : Directed (· ⊆ ·) s) (h : ∀ i, sSupIndep (s i)) : + sSupIndep (⋃ i, s i) := by by_cases hη : Nonempty η - · rw [CompleteLattice.setIndependent_iff_finite] + · rw [sSupIndep_iff_finite] intro t ht obtain ⟨I, fi, hI⟩ := Set.finite_subset_iUnion t.finite_toSet ht obtain ⟨i, hi⟩ := hs.finset_le fi.toFinset @@ -465,10 +490,16 @@ theorem CompleteLattice.setIndependent_iUnion_of_directed {η : Type*} {s : η exfalso exact hη ⟨i⟩ -theorem CompleteLattice.independent_sUnion_of_directed {s : Set (Set α)} (hs : DirectedOn (· ⊆ ·) s) - (h : ∀ a ∈ s, CompleteLattice.SetIndependent a) : CompleteLattice.SetIndependent (⋃₀ s) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_iUnion_of_directed := sSupIndep_iUnion_of_directed + +theorem iSupIndep_sUnion_of_directed {s : Set (Set α)} (hs : DirectedOn (· ⊆ ·) s) + (h : ∀ a ∈ s, sSupIndep a) : sSupIndep (⋃₀ s) := by rw [Set.sUnion_eq_iUnion] - exact CompleteLattice.setIndependent_iUnion_of_directed hs.directed_val (by simpa using h) + exact sSupIndep_iUnion_of_directed hs.directed_val (by simpa using h) + +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_sUnion_of_directed := iSupIndep_sUnion_of_directed end @@ -495,11 +526,11 @@ alias IsSupFiniteCompact.wellFounded := IsSupFiniteCompact.wellFoundedGT @[deprecated (since := "2024-10-07")] alias _root_.WellFounded.isSupClosedCompact := WellFoundedGT.isSupClosedCompact @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_of_setIndependent := WellFoundedGT.finite_of_setIndependent +alias WellFounded.finite_of_setIndependent := WellFoundedGT.finite_of_sSupIndep @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_ne_bot_of_independent := WellFoundedGT.finite_ne_bot_of_independent +alias WellFounded.finite_ne_bot_of_independent := WellFoundedGT.finite_ne_bot_of_iSupIndep @[deprecated (since := "2024-10-07")] -alias WellFounded.finite_of_independent := WellFoundedGT.finite_of_independent +alias WellFounded.finite_of_independent := WellFoundedGT.finite_of_iSupIndep @[deprecated (since := "2024-10-07")] alias isCompactlyGenerated_of_wellFounded := isCompactlyGenerated_of_wellFoundedGT @@ -578,16 +609,16 @@ Most explicitly, every element is the complement of a supremum of indepedendent /-- In an atomic lattice, every element `b` has a complement of the form `sSup s`, where each element of `s` is an atom. See also `complementedLattice_of_sSup_atoms_eq_top`. -/ -theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } = ⊤) (b : α) : - ∃ s : Set α, CompleteLattice.SetIndependent s ∧ +theorem exists_sSupIndep_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } = ⊤) (b : α) : + ∃ s : Set α, sSupIndep s ∧ IsCompl b (sSup s) ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := by -- porting note(https://github.com/leanprover-community/mathlib4/issues/5732): -- `obtain` chokes on the placeholder. have zorn := zorn_subset - (S := {s : Set α | CompleteLattice.SetIndependent s ∧ Disjoint b (sSup s) ∧ ∀ a ∈ s, IsAtom a}) + (S := {s : Set α | sSupIndep s ∧ Disjoint b (sSup s) ∧ ∀ a ∈ s, IsAtom a}) fun c hc1 hc2 => ⟨⋃₀ c, - ⟨CompleteLattice.independent_sUnion_of_directed hc2.directedOn fun s hs => (hc1 hs).1, ?_, + ⟨iSupIndep_sUnion_of_directed hc2.directedOn fun s hs => (hc1 hs).1, ?_, fun a ⟨s, sc, as⟩ => (hc1 sc).2.2 a as⟩, fun _ => Set.subset_sUnion_of_mem⟩ swap @@ -633,16 +664,22 @@ theorem exists_setIndependent_isCompl_sSup_atoms (h : sSup { a : α | IsAtom a } · exact s_atoms x hx · exact ha -theorem exists_setIndependent_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : - ∃ s : Set α, CompleteLattice.SetIndependent s ∧ sSup s = ⊤ ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := - let ⟨s, s_ind, s_top, s_atoms⟩ := exists_setIndependent_isCompl_sSup_atoms h ⊥ +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_isCompl_sSup_atoms := exists_sSupIndep_isCompl_sSup_atoms + +theorem exists_sSupIndep_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : + ∃ s : Set α, sSupIndep s ∧ sSup s = ⊤ ∧ ∀ ⦃a⦄, a ∈ s → IsAtom a := + let ⟨s, s_ind, s_top, s_atoms⟩ := exists_sSupIndep_isCompl_sSup_atoms h ⊥ ⟨s, s_ind, eq_top_of_isCompl_bot s_top.symm, s_atoms⟩ +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_of_sSup_atoms_eq_top := exists_sSupIndep_of_sSup_atoms_eq_top + /-- See [Theorem 6.6][calugareanu]. -/ theorem complementedLattice_of_sSup_atoms_eq_top (h : sSup { a : α | IsAtom a } = ⊤) : ComplementedLattice α := ⟨fun b => - let ⟨s, _, s_top, _⟩ := exists_setIndependent_isCompl_sSup_atoms h b + let ⟨s, _, s_top, _⟩ := exists_sSupIndep_isCompl_sSup_atoms h b ⟨sSup s, s_top⟩⟩ /-- See [Theorem 6.6][calugareanu]. -/ diff --git a/Mathlib/Order/SupIndep.lean b/Mathlib/Order/SupIndep.lean index 68e516ec2adec..27770ea7eaa4a 100644 --- a/Mathlib/Order/SupIndep.lean +++ b/Mathlib/Order/SupIndep.lean @@ -18,19 +18,19 @@ sup-independent if, for all `a`, `f a` and the supremum of the rest are disjoint ## Main definitions * `Finset.SupIndep s f`: a family of elements `f` are supremum independent on the finite set `s`. -* `CompleteLattice.SetIndependent s`: a set of elements are supremum independent. -* `CompleteLattice.Independent f`: a family of elements are supremum independent. +* `sSupIndep s`: a set of elements are supremum independent. +* `iSupIndep f`: a family of elements are supremum independent. ## Main statements * In a distributive lattice, supremum independence is equivalent to pairwise disjointness: * `Finset.supIndep_iff_pairwiseDisjoint` - * `CompleteLattice.setIndependent_iff_pairwiseDisjoint` - * `CompleteLattice.independent_iff_pairwiseDisjoint` + * `CompleteLattice.sSupIndep_iff_pairwiseDisjoint` + * `CompleteLattice.iSupIndep_iff_pairwiseDisjoint` * Otherwise, supremum independence is stronger than pairwise disjointness: * `Finset.SupIndep.pairwiseDisjoint` - * `CompleteLattice.SetIndependent.pairwiseDisjoint` - * `CompleteLattice.Independent.pairwiseDisjoint` + * `sSupIndep.pairwiseDisjoint` + * `iSupIndep.pairwiseDisjoint` ## Implementation notes @@ -88,7 +88,7 @@ theorem SupIndep.le_sup_iff (hs : s.SupIndep f) (hts : t ⊆ s) (hi : i ∈ s) ( by_contra hit exact hf i (disjoint_self.1 <| (hs hts hi hit).mono_right h) -/-- The RHS looks like the definition of `CompleteLattice.Independent`. -/ +/-- The RHS looks like the definition of `iSupIndep`. -/ theorem supIndep_iff_disjoint_erase [DecidableEq ι] : s.SupIndep f ↔ ∀ i ∈ s, Disjoint (f i) ((s.erase i).sup f) := ⟨fun hs _ hi => hs (erase_subset _ _) hi (not_mem_erase _ _), fun hs _ ht i hi hit => @@ -269,38 +269,48 @@ end Finset /-! ### On complete lattices via `sSup` -/ - -namespace CompleteLattice - +section CompleteLattice variable [CompleteLattice α] open Set Function /-- An independent set of elements in a complete lattice is one in which every element is disjoint from the `Sup` of the rest. -/ -def SetIndependent (s : Set α) : Prop := +def sSupIndep (s : Set α) : Prop := ∀ ⦃a⦄, a ∈ s → Disjoint a (sSup (s \ {a})) -variable {s : Set α} (hs : SetIndependent s) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.SetIndependent := sSupIndep + +variable {s : Set α} (hs : sSupIndep s) @[simp] -theorem setIndependent_empty : SetIndependent (∅ : Set α) := fun x hx => +theorem sSupIndep_empty : sSupIndep (∅ : Set α) := fun x hx => (Set.not_mem_empty x hx).elim +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_empty := sSupIndep_empty + include hs in -theorem SetIndependent.mono {t : Set α} (hst : t ⊆ s) : SetIndependent t := fun _ ha => +theorem sSupIndep.mono {t : Set α} (hst : t ⊆ s) : sSupIndep t := fun _ ha => (hs (hst ha)).mono_right (sSup_le_sSup (diff_subset_diff_left hst)) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.SetIndependent.mono := sSupIndep.mono + include hs in /-- If the elements of a set are independent, then any pair within that set is disjoint. -/ -theorem SetIndependent.pairwiseDisjoint : s.PairwiseDisjoint id := fun _ hx y hy h => +theorem sSupIndep.pairwiseDisjoint : s.PairwiseDisjoint id := fun _ hx y hy h => disjoint_sSup_right (hs hx) ((mem_diff y).mpr ⟨hy, h.symm⟩) -theorem setIndependent_singleton (a : α) : SetIndependent ({a} : Set α) := fun i hi ↦ by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.SetIndependent.pairwiseDisjoint := sSupIndep.pairwiseDisjoint + +theorem sSupIndep_singleton (a : α) : sSupIndep ({a} : Set α) := fun i hi ↦ by simp_all -theorem setIndependent_pair {a b : α} (hab : a ≠ b) : - SetIndependent ({a, b} : Set α) ↔ Disjoint a b := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.setIndependent_singleton := sSupIndep_singleton + +theorem sSupIndep_pair {a b : α} (hab : a ≠ b) : + sSupIndep ({a, b} : Set α) ↔ Disjoint a b := by constructor · intro h exact h.pairwiseDisjoint (mem_insert _ _) (mem_insert_of_mem _ (mem_singleton _)) hab @@ -310,15 +320,20 @@ theorem setIndependent_pair {a b : α} (hab : a ≠ b) : · convert h.symm using 1 simp [hab, sSup_singleton] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_pair := sSupIndep_pair + include hs in /-- If the elements of a set are independent, then any element is disjoint from the `sSup` of some subset of the rest. -/ -theorem SetIndependent.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : y ⊆ s) (hxy : x ∉ y) : +theorem sSupIndep.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : y ⊆ s) (hxy : x ∉ y) : Disjoint x (sSup y) := by have := (hs.mono <| insert_subset_iff.mpr ⟨hx, hy⟩) (mem_insert x _) rw [insert_diff_of_mem _ (mem_singleton _), diff_singleton_eq_self hxy] at this exact this +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.SetIndependent.disjoint_sSup := sSupIndep.disjoint_sSup + /-- An independent indexed family of elements in a complete lattice is one in which every element is disjoint from the `iSup` of the rest. @@ -328,70 +343,96 @@ theorem SetIndependent.disjoint_sSup {x : α} {y : Set α} (hx : x ∈ s) (hy : Example: an indexed family of submodules of a module is independent in this sense if and only the natural map from the direct sum of the submodules to the module is injective. -/ -def Independent {ι : Sort*} {α : Type*} [CompleteLattice α] (t : ι → α) : Prop := +def iSupIndep {ι : Sort*} {α : Type*} [CompleteLattice α] (t : ι → α) : Prop := ∀ i : ι, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) -theorem setIndependent_iff {α : Type*} [CompleteLattice α] (s : Set α) : - SetIndependent s ↔ Independent ((↑) : s → α) := by - simp_rw [Independent, SetIndependent, SetCoe.forall, sSup_eq_iSup] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent := iSupIndep + +theorem sSupIndep_iff {α : Type*} [CompleteLattice α] (s : Set α) : + sSupIndep s ↔ iSupIndep ((↑) : s → α) := by + simp_rw [iSupIndep, sSupIndep, SetCoe.forall, sSup_eq_iSup] refine forall₂_congr fun a ha => ?_ simp [iSup_subtype, iSup_and] -variable {t : ι → α} (ht : Independent t) +@[deprecated (since := "2024-11-24")] alias CompleteLattice.setIndependent_iff := sSupIndep_iff + +variable {t : ι → α} (ht : iSupIndep t) -theorem independent_def : Independent t ↔ ∀ i : ι, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) := +theorem iSupIndep_def : iSupIndep t ↔ ∀ i, Disjoint (t i) (⨆ (j) (_ : j ≠ i), t j) := Iff.rfl -theorem independent_def' : Independent t ↔ ∀ i, Disjoint (t i) (sSup (t '' { j | j ≠ i })) := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def := iSupIndep_def + +theorem iSupIndep_def' : iSupIndep t ↔ ∀ i, Disjoint (t i) (sSup (t '' { j | j ≠ i })) := by simp_rw [sSup_image] rfl -theorem independent_def'' : - Independent t ↔ ∀ i, Disjoint (t i) (sSup { a | ∃ j ≠ i, t j = a }) := by - rw [independent_def'] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def' := iSupIndep_def' + +theorem iSupIndep_def'' : + iSupIndep t ↔ ∀ i, Disjoint (t i) (sSup { a | ∃ j ≠ i, t j = a }) := by + rw [iSupIndep_def'] aesop +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_def'' := iSupIndep_def'' + @[simp] -theorem independent_empty (t : Empty → α) : Independent t := +theorem iSupIndep_empty (t : Empty → α) : iSupIndep t := nofun +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_empty := iSupIndep_empty + @[simp] -theorem independent_pempty (t : PEmpty → α) : Independent t := +theorem iSupIndep_pempty (t : PEmpty → α) : iSupIndep t := nofun +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_pempty := iSupIndep_pempty + include ht in /-- If the elements of a set are independent, then any pair within that set is disjoint. -/ -theorem Independent.pairwiseDisjoint : Pairwise (Disjoint on t) := fun x y h => +theorem iSupIndep.pairwiseDisjoint : Pairwise (Disjoint on t) := fun x y h => disjoint_sSup_right (ht x) ⟨y, iSup_pos h.symm⟩ -theorem Independent.mono {s t : ι → α} (hs : Independent s) (hst : t ≤ s) : Independent t := +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.pairwiseDisjoint := iSupIndep.pairwiseDisjoint + +theorem iSupIndep.mono {s t : ι → α} (hs : iSupIndep s) (hst : t ≤ s) : iSupIndep t := fun i => (hs i).mono (hst i) <| iSup₂_mono fun j _ => hst j +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.mono := iSupIndep.mono + /-- Composing an independent indexed family with an injective function on the index results in another indepedendent indexed family. -/ -theorem Independent.comp {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : Independent t) - (hf : Injective f) : Independent (t ∘ f) := fun i => +theorem iSupIndep.comp {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : iSupIndep t) + (hf : Injective f) : iSupIndep (t ∘ f) := fun i => (ht (f i)).mono_right <| by refine (iSup_mono fun i => ?_).trans (iSup_comp_le _ f) exact iSup_const_mono hf.ne -theorem Independent.comp' {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : Independent <| t ∘ f) - (hf : Surjective f) : Independent t := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.comp := iSupIndep.comp + +theorem iSupIndep.comp' {ι ι' : Sort*} {t : ι → α} {f : ι' → ι} (ht : iSupIndep <| t ∘ f) + (hf : Surjective f) : iSupIndep t := by intro i obtain ⟨i', rfl⟩ := hf i rw [← hf.iSup_comp] exact (ht i').mono_right (biSup_mono fun j' hij => mt (congr_arg f) hij) -theorem Independent.setIndependent_range (ht : Independent t) : SetIndependent <| range t := by - rw [setIndependent_iff] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.comp' := iSupIndep.comp' + +theorem iSupIndep.sSupIndep_range (ht : iSupIndep t) : sSupIndep <| range t := by + rw [sSupIndep_iff] rw [← coe_comp_rangeFactorization t] at ht exact ht.comp' surjective_onto_range +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.setIndependent_range := iSupIndep.sSupIndep_range + @[simp] -theorem independent_ne_bot_iff_independent : - Independent (fun i : {i // t i ≠ ⊥} ↦ t i) ↔ Independent t := by +theorem iSupIndep_ne_bot : + iSupIndep (fun i : {i // t i ≠ ⊥} ↦ t i) ↔ iSupIndep t := by refine ⟨fun h ↦ ?_, fun h ↦ h.comp Subtype.val_injective⟩ - simp only [independent_def] at h ⊢ + simp only [iSupIndep_def] at h ⊢ intro i cases eq_or_ne (t i) ⊥ with | inl hi => simp [hi] @@ -402,7 +443,10 @@ theorem independent_ne_bot_iff_independent : simp only [iSup_comm (ι' := _ ≠ i), this, ne_eq, sup_of_le_right, Subtype.mk.injEq, iSup_bot, bot_le] -theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_ne_bot_iff_independent := iSupIndep_ne_bot + +theorem iSupIndep.injOn (ht : iSupIndep t) : InjOn t {i | t i ≠ ⊥} := by rintro i _ j (hj : t j ≠ ⊥) h by_contra! contra apply hj @@ -413,12 +457,17 @@ theorem Independent.injOn (ht : Independent t) : InjOn t {i | t i ≠ ⊥} := by -- Porting note: needs explicit `f` exact le_iSup₂ (f := fun x _ ↦ t x) j contra -theorem Independent.injective (ht : Independent t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by +@[deprecated (since := "2024-11-24")] alias CompleteLattice.Independent.injOn := iSupIndep.injOn + +theorem iSupIndep.injective (ht : iSupIndep t) (h_ne_bot : ∀ i, t i ≠ ⊥) : Injective t := by suffices univ = {i | t i ≠ ⊥} by rw [injective_iff_injOn_univ, this]; exact ht.injOn aesop -theorem independent_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k = j) : - Independent t ↔ Disjoint (t i) (t j) := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.injective := iSupIndep.injective + +theorem iSupIndep_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k = j) : + iSupIndep t ↔ Disjoint (t i) (t j) := by constructor · exact fun h => h.pairwiseDisjoint hij · rintro h k @@ -428,37 +477,49 @@ theorem independent_pair {i j : ι} (hij : i ≠ j) (huniv : ∀ k, k = i ∨ k · refine h.symm.mono_right (iSup_le fun j => iSup_le fun hj => Eq.le ?_) rw [(huniv j).resolve_right hj] +@[deprecated (since := "2024-11-24")] alias CompleteLattice.independent_pair := iSupIndep_pair + /-- Composing an independent indexed family with an order isomorphism on the elements results in another independent indexed family. -/ -theorem Independent.map_orderIso {ι : Sort*} {α β : Type*} [CompleteLattice α] - [CompleteLattice β] (f : α ≃o β) {a : ι → α} (ha : Independent a) : Independent (f ∘ a) := +theorem iSupIndep.map_orderIso {ι : Sort*} {α β : Type*} [CompleteLattice α] + [CompleteLattice β] (f : α ≃o β) {a : ι → α} (ha : iSupIndep a) : iSupIndep (f ∘ a) := fun i => ((ha i).map_orderIso f).mono_right (f.monotone.le_map_iSup₂ _) +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.map_orderIso := iSupIndep.map_orderIso + @[simp] -theorem independent_map_orderIso_iff {ι : Sort*} {α β : Type*} [CompleteLattice α] - [CompleteLattice β] (f : α ≃o β) {a : ι → α} : Independent (f ∘ a) ↔ Independent a := +theorem iSupIndep_map_orderIso_iff {ι : Sort*} {α β : Type*} [CompleteLattice α] + [CompleteLattice β] (f : α ≃o β) {a : ι → α} : iSupIndep (f ∘ a) ↔ iSupIndep a := ⟨fun h => have hf : f.symm ∘ f ∘ a = a := congr_arg (· ∘ a) f.left_inv.comp_eq_id hf ▸ h.map_orderIso f.symm, fun h => h.map_orderIso f⟩ +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_map_orderIso_iff := iSupIndep_map_orderIso_iff + /-- If the elements of a set are independent, then any element is disjoint from the `iSup` of some subset of the rest. -/ -theorem Independent.disjoint_biSup {ι : Type*} {α : Type*} [CompleteLattice α] {t : ι → α} - (ht : Independent t) {x : ι} {y : Set ι} (hx : x ∉ y) : Disjoint (t x) (⨆ i ∈ y, t i) := +theorem iSupIndep.disjoint_biSup {ι : Type*} {α : Type*} [CompleteLattice α] {t : ι → α} + (ht : iSupIndep t) {x : ι} {y : Set ι} (hx : x ∉ y) : Disjoint (t x) (⨆ i ∈ y, t i) := Disjoint.mono_right (biSup_mono fun _ hi => (ne_of_mem_of_not_mem hi hx : _)) (ht x) -lemma independent_of_independent_coe_Iic_comp {ι : Sort*} {a : α} {t : ι → Set.Iic a} - (ht : Independent ((↑) ∘ t : ι → α)) : Independent t := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.disjoint_biSup := iSupIndep.disjoint_biSup + +lemma iSupIndep.of_coe_Iic_comp {ι : Sort*} {a : α} {t : ι → Set.Iic a} + (ht : iSupIndep ((↑) ∘ t : ι → α)) : iSupIndep t := by intro i x specialize ht i simp_rw [Function.comp_apply, ← Set.Iic.coe_iSup] at ht exact @ht x -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_of_independent_coe_Iic_comp := iSupIndep.of_coe_Iic_comp -theorem CompleteLattice.independent_iff_supIndep [CompleteLattice α] {s : Finset ι} {f : ι → α} : - CompleteLattice.Independent (f ∘ ((↑) : s → ι)) ↔ s.SupIndep f := by +theorem iSupIndep_iff_supIndep {s : Finset ι} {f : ι → α} : + iSupIndep (f ∘ ((↑) : s → ι)) ↔ s.SupIndep f := by classical rw [Finset.supIndep_iff_disjoint_erase] refine Subtype.forall.trans (forall₂_congr fun a b => ?_) @@ -468,39 +529,47 @@ theorem CompleteLattice.independent_iff_supIndep [CompleteLattice α] {s : Finse congr! 1 simp [iSup_and, @iSup_comm _ (_ ∈ s)] -alias ⟨CompleteLattice.Independent.supIndep, Finset.SupIndep.independent⟩ := - CompleteLattice.independent_iff_supIndep +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep := iSupIndep_iff_supIndep + +alias ⟨iSupIndep.supIndep, Finset.SupIndep.independent⟩ := iSupIndep_iff_supIndep -theorem CompleteLattice.Independent.supIndep' [CompleteLattice α] {f : ι → α} (s : Finset ι) - (h : CompleteLattice.Independent f) : s.SupIndep f := - CompleteLattice.Independent.supIndep (h.comp Subtype.coe_injective) +theorem iSupIndep.supIndep' {f : ι → α} (s : Finset ι) (h : iSupIndep f) : s.SupIndep f := + iSupIndep.supIndep (h.comp Subtype.coe_injective) -/-- A variant of `CompleteLattice.independent_iff_supIndep` for `Fintype`s. -/ -theorem CompleteLattice.independent_iff_supIndep_univ [CompleteLattice α] [Fintype ι] {f : ι → α} : - CompleteLattice.Independent f ↔ Finset.univ.SupIndep f := by +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.Independent.supIndep' := iSupIndep.supIndep' + +/-- A variant of `CompleteLattice.iSupIndep_iff_supIndep` for `Fintype`s. -/ +theorem iSupIndep_iff_supIndep_univ [Fintype ι] {f : ι → α} : + iSupIndep f ↔ Finset.univ.SupIndep f := by classical - simp [Finset.supIndep_iff_disjoint_erase, CompleteLattice.Independent, Finset.sup_eq_iSup] + simp [Finset.supIndep_iff_disjoint_erase, iSupIndep, Finset.sup_eq_iSup] -alias ⟨CompleteLattice.Independent.sup_indep_univ, Finset.SupIndep.independent_of_univ⟩ := - CompleteLattice.independent_iff_supIndep_univ +@[deprecated (since := "2024-11-24")] +alias CompleteLattice.independent_iff_supIndep_univ := iSupIndep_iff_supIndep_univ -section Frame +alias ⟨iSupIndep.sup_indep_univ, Finset.SupIndep.iSupIndep_of_univ⟩ := iSupIndep_iff_supIndep_univ -namespace CompleteLattice +end CompleteLattice +section Frame variable [Order.Frame α] -theorem setIndependent_iff_pairwiseDisjoint {s : Set α} : - SetIndependent s ↔ s.PairwiseDisjoint id := - ⟨SetIndependent.pairwiseDisjoint, fun hs _ hi => +theorem sSupIndep_iff_pairwiseDisjoint {s : Set α} : sSupIndep s ↔ s.PairwiseDisjoint id := + ⟨sSupIndep.pairwiseDisjoint, fun hs _ hi => disjoint_sSup_iff.2 fun _ hj => hs hi hj.1 <| Ne.symm hj.2⟩ -alias ⟨_, _root_.Set.PairwiseDisjoint.setIndependent⟩ := setIndependent_iff_pairwiseDisjoint +@[deprecated (since := "2024-11-24")] +alias setIndependent_iff_pairwiseDisjoint := sSupIndep_iff_pairwiseDisjoint -theorem independent_iff_pairwiseDisjoint {f : ι → α} : Independent f ↔ Pairwise (Disjoint on f) := - ⟨Independent.pairwiseDisjoint, fun hs _ => +alias ⟨_, _root_.Set.PairwiseDisjoint.sSupIndep⟩ := sSupIndep_iff_pairwiseDisjoint + +theorem iSupIndep_iff_pairwiseDisjoint {f : ι → α} : iSupIndep f ↔ Pairwise (Disjoint on f) := + ⟨iSupIndep.pairwiseDisjoint, fun hs _ => disjoint_iSup_iff.2 fun _ => disjoint_iSup_iff.2 fun hij => hs hij.symm⟩ -end CompleteLattice +@[deprecated (since := "2024-11-24")] +alias independent_iff_pairwiseDisjoint := iSupIndep_iff_pairwiseDisjoint end Frame diff --git a/Mathlib/RingTheory/FiniteLength.lean b/Mathlib/RingTheory/FiniteLength.lean index b1cf98eb550b8..8eea41b4e20e2 100644 --- a/Mathlib/RingTheory/FiniteLength.lean +++ b/Mathlib/RingTheory/FiniteLength.lean @@ -78,13 +78,13 @@ theorem isFiniteLength_iff_exists_compositionSeries : theorem IsSemisimpleModule.finite_tfae [IsSemisimpleModule R M] : List.TFAE [Module.Finite R M, IsNoetherian R M, IsArtinian R M, IsFiniteLength R M, - ∃ s : Set (Submodule R M), s.Finite ∧ CompleteLattice.SetIndependent s ∧ + ∃ s : Set (Submodule R M), s.Finite ∧ sSupIndep s ∧ sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m] := by rw [isFiniteLength_iff_isNoetherian_isArtinian] - obtain ⟨s, hs⟩ := IsSemisimpleModule.exists_setIndependent_sSup_simples_eq_top R M + obtain ⟨s, hs⟩ := IsSemisimpleModule.exists_sSupIndep_sSup_simples_eq_top R M tfae_have 1 ↔ 2 := ⟨fun _ ↦ inferInstance, fun _ ↦ inferInstance⟩ - tfae_have 2 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedGT.finite_of_setIndependent hs.1, hs⟩ - tfae_have 3 → 5 := fun _ ↦ ⟨s, CompleteLattice.WellFoundedLT.finite_of_setIndependent hs.1, hs⟩ + tfae_have 2 → 5 := fun _ ↦ ⟨s, WellFoundedGT.finite_of_sSupIndep hs.1, hs⟩ + tfae_have 3 → 5 := fun _ ↦ ⟨s, WellFoundedLT.finite_of_sSupIndep hs.1, hs⟩ tfae_have 5 → 4 := fun ⟨s, fin, _, sSup_eq_top, simple⟩ ↦ by rw [← isNoetherian_top_iff, ← Submodule.topEquiv.isArtinian_iff, ← sSup_eq_top, sSup_eq_iSup, ← iSup_subtype''] diff --git a/Mathlib/RingTheory/Noetherian/Defs.lean b/Mathlib/RingTheory/Noetherian/Defs.lean index 953d4b7265411..926f9ba8c6880 100644 --- a/Mathlib/RingTheory/Noetherian/Defs.lean +++ b/Mathlib/RingTheory/Noetherian/Defs.lean @@ -313,17 +313,20 @@ universe w variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] -lemma Submodule.finite_ne_bot_of_independent {ι : Type*} {N : ι → Submodule R M} - (h : CompleteLattice.Independent N) : +lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} + (h : iSupIndep N) : Set.Finite {i | N i ≠ ⊥} := - CompleteLattice.WellFoundedGT.finite_ne_bot_of_independent h + WellFoundedGT.finite_ne_bot_of_iSupIndep h + +@[deprecated (since := "2024-11-24")] +alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep /-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if the module is Noetherian. -/ theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} (hv : LinearIndependent R v) : Finite ι := by - refine CompleteLattice.WellFoundedGT.finite_of_independent - hv.independent_span_singleton + refine WellFoundedGT.finite_of_iSupIndep + hv.iSupIndep_span_singleton fun i contra => ?_ apply hv.ne_zero i have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) diff --git a/Mathlib/RingTheory/SimpleModule.lean b/Mathlib/RingTheory/SimpleModule.lean index 66282ba413b64..dbae9bcb743ed 100644 --- a/Mathlib/RingTheory/SimpleModule.lean +++ b/Mathlib/RingTheory/SimpleModule.lean @@ -198,12 +198,14 @@ theorem exists_simple_submodule [Nontrivial M] : ∃ m : Submodule R M, IsSimple theorem sSup_simples_eq_top : sSup { m : Submodule R M | IsSimpleModule R m } = ⊤ := by simpa only [isSimpleModule_iff_isAtom] using sSup_atoms_eq_top -theorem exists_setIndependent_sSup_simples_eq_top : - ∃ s : Set (Submodule R M), CompleteLattice.SetIndependent s ∧ - sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m := by +theorem exists_sSupIndep_sSup_simples_eq_top : + ∃ s : Set (Submodule R M), sSupIndep s ∧ sSup s = ⊤ ∧ ∀ m ∈ s, IsSimpleModule R m := by have := sSup_simples_eq_top R M simp_rw [isSimpleModule_iff_isAtom] at this ⊢ - exact exists_setIndependent_of_sSup_atoms_eq_top this + exact exists_sSupIndep_of_sSup_atoms_eq_top this + +@[deprecated (since := "2024-11-24")] +alias exists_setIndependent_sSup_simples_eq_top := exists_sSupIndep_sSup_simples_eq_top /-- The annihilator of a semisimple module over a commutative ring is a radical ideal. -/ theorem annihilator_isRadical (R) [CommRing R] [Module R M] [IsSemisimpleModule R M] : diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index fd905374bfebb..a5ca83037daf6 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -686,11 +686,11 @@ CompactIccSpace.mk' CompactIccSpace.mk'' CompHaus.toProfinite_obj' compl_beattySeq' -CompleteLattice.Independent.comp' -CompleteLattice.independent_def' -CompleteLattice.independent_def'' -CompleteLattice.independent_of_dfinsupp_sumAddHom_injective' -CompleteLattice.Independent.supIndep' +iSupIndep.comp' +iSupIndep_def' +iSupIndep_def'' +iSupIndep_of_dfinsupp_sumAddHom_injective' +iSupIndep.supIndep' CompleteLattice.inf_continuous' CompleteLattice.sSup_continuous' CompletelyDistribLattice.MinimalAxioms.iInf_iSup_eq' @@ -2233,7 +2233,7 @@ LieIdeal.map_sup_ker_eq_map' LieModule.chainTop_isNonZero' LieModule.coe_chainTop' LieModule.genWeightSpaceChain_def' -LieModule.independent_genWeightSpace' +LieModule.iSupIndep_genWeightSpace' LieModule.instIsTrivialOfSubsingleton' LieModule.isNilpotent_of_top_iff' LieModule.iSup_genWeightSpace_eq_top' From b96544419e0739c6aefecfc4ee368400a77266ce Mon Sep 17 00:00:00 2001 From: Eric Wieser Date: Mon, 25 Nov 2024 10:13:40 +0000 Subject: [PATCH 129/250] doc: describe why map_ofNat can't be simp (yet) (#19410) See the benchmark in #19388, which while not disastrous is not ideal either. Unless the situation changes with https://github.com/leanprover/lean4/issues/2867, the DiscrTree key would apply to all funlike operations applied to any argument, not just applied to numerals. --- Mathlib/Data/Nat/Cast/Basic.lean | 5 +++++ 1 file changed, 5 insertions(+) diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 63025e62d2769..c494b68c05067 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -159,6 +159,11 @@ theorem eq_natCast [FunLike F ℕ R] [RingHomClass F ℕ R] (f : F) : ∀ n, f n theorem map_natCast [FunLike F R S] [RingHomClass F R S] (f : F) : ∀ n : ℕ, f (n : R) = n := map_natCast' f <| map_one f +/-- This lemma is not marked `@[simp]` lemma because its `#discr_tree_key` (for the LHS) would just +be `DFunLike.coe _ _`, due to the `no_index` that https://github.com/leanprover/lean4/issues/2867 +forces us to include, and therefore it would negatively impact performance. + +If that issue is resolved, this can be marked `@[simp]`. -/ theorem map_ofNat [FunLike F R S] [RingHomClass F R S] (f : F) (n : ℕ) [Nat.AtLeastTwo n] : (f (no_index (OfNat.ofNat n)) : S) = OfNat.ofNat n := map_natCast f n From 552d98cad0c88aa12d997640f911c0059541a564 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 21:39:19 +1100 Subject: [PATCH 130/250] bump batteries --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index cdc932a7c6e5d..37d601960eb6e 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0681948d281fb8d3499952cd954be673c17e6e81", + "rev": "332e0aa7f17c64a29dded0c373a7940600069e5b", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From 040087e5c7622b5ae901a50c863eb0bb8076d935 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 21:43:39 +1100 Subject: [PATCH 131/250] disambiguate Vector --- Mathlib/Computability/Primrec.lean | 38 +++++++++-------- Mathlib/Computability/TuringMachine.lean | 9 ++-- Mathlib/Data/Fintype/BigOperators.lean | 3 +- Mathlib/Data/Fintype/Fin.lean | 2 +- Mathlib/Data/Fintype/Vector.lean | 2 +- Mathlib/Data/Num/Bitwise.lean | 2 +- Mathlib/Data/Sym/Basic.lean | 15 ++++--- Mathlib/Data/Sym/Sym2.lean | 2 +- Mathlib/Logic/Equiv/List.lean | 4 +- Mathlib/Logic/Small/List.lean | 4 +- MathlibTest/fractran.lean | 54 ++++++++++++++++++++++++ 11 files changed, 97 insertions(+), 38 deletions(-) create mode 100644 MathlibTest/fractran.lean diff --git a/Mathlib/Computability/Primrec.lean b/Mathlib/Computability/Primrec.lean index f7492a1e781aa..76aceca0eb52f 100644 --- a/Mathlib/Computability/Primrec.lean +++ b/Mathlib/Computability/Primrec.lean @@ -1101,7 +1101,7 @@ def subtype {p : α → Prop} [DecidablePred p] (hp : PrimrecPred p) : Primcodab instance fin {n} : Primcodable (Fin n) := @ofEquiv _ _ (subtype <| nat_lt.comp .id (const n)) Fin.equivSubtype -instance vector {n} : Primcodable (Vector α n) := +instance vector {n} : Primcodable (Mathlib.Vector α n) := subtype ((@Primrec.eq ℕ _ _).comp list_length (const _)) instance finArrow {n} : Primcodable (Fin n → α) := @@ -1184,25 +1184,26 @@ theorem fin_val {n} : Primrec (fun (i : Fin n) => (i : ℕ)) := theorem fin_succ {n} : Primrec (@Fin.succ n) := fin_val_iff.1 <| by simp [succ.comp fin_val] -theorem vector_toList {n} : Primrec (@Vector.toList α n) := +theorem vector_toList {n} : Primrec (@Mathlib.Vector.toList α n) := subtype_val -theorem vector_toList_iff {n} {f : α → Vector β n} : (Primrec fun a => (f a).toList) ↔ Primrec f := +theorem vector_toList_iff {n} {f : α → Mathlib.Vector β n} : + (Primrec fun a => (f a).toList) ↔ Primrec f := subtype_val_iff -theorem vector_cons {n} : Primrec₂ (@Vector.cons α n) := +theorem vector_cons {n} : Primrec₂ (@Mathlib.Vector.cons α n) := vector_toList_iff.1 <| by simpa using list_cons.comp fst (vector_toList_iff.2 snd) -theorem vector_length {n} : Primrec (@Vector.length α n) := +theorem vector_length {n} : Primrec (@Mathlib.Vector.length α n) := const _ -theorem vector_head {n} : Primrec (@Vector.head α n) := +theorem vector_head {n} : Primrec (@Mathlib.Vector.head α n) := option_some_iff.1 <| (list_head?.comp vector_toList).of_eq fun ⟨_ :: _, _⟩ => rfl -theorem vector_tail {n} : Primrec (@Vector.tail α n) := +theorem vector_tail {n} : Primrec (@Mathlib.Vector.tail α n) := vector_toList_iff.1 <| (list_tail.comp vector_toList).of_eq fun ⟨l, h⟩ => by cases l <;> rfl -theorem vector_get {n} : Primrec₂ (@Vector.get α n) := +theorem vector_get {n} : Primrec₂ (@Mathlib.Vector.get α n) := option_some_iff.1 <| (list_get?.comp (vector_toList.comp fst) (fin_val.comp snd)).of_eq fun a => by rw [Vector.get_eq_get, ← List.get?_eq_get] @@ -1215,13 +1216,13 @@ theorem list_ofFn : simpa [List.ofFn_succ] using list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Primrec (f i)) : - Primrec fun a => Vector.ofFn fun i => f i a := + Primrec fun a => Mathlib.Vector.ofFn fun i => f i a := vector_toList_iff.1 <| by simp [list_ofFn hf] -theorem vector_get' {n} : Primrec (@Vector.get α n) := +theorem vector_get' {n} : Primrec (@Mathlib.Vector.get α n) := of_equiv_symm -theorem vector_ofFn' {n} : Primrec (@Vector.ofFn α n) := +theorem vector_ofFn' {n} : Primrec (@Mathlib.Vector.ofFn α n) := of_equiv theorem fin_app {n} : Primrec₂ (@id (Fin n → σ)) := @@ -1248,16 +1249,16 @@ open Mathlib.Vector work with n-ary functions on ℕ instead of unary functions. We prove that this is equivalent to the regular notion in `to_prim` and `of_prim`. -/ -inductive Primrec' : ∀ {n}, (Vector ℕ n → ℕ) → Prop +inductive Primrec' : ∀ {n}, (Mathlib.Vector ℕ n → ℕ) → Prop | zero : @Primrec' 0 fun _ => 0 | succ : @Primrec' 1 fun v => succ v.head | get {n} (i : Fin n) : Primrec' fun v => v.get i - | comp {m n f} (g : Fin n → Vector ℕ m → ℕ) : - Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (ofFn fun i => g i a) + | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m → ℕ) : + Primrec' f → (∀ i, Primrec' (g i)) → Primrec' fun a => f (Mathlib.Vector.ofFn fun i => g i a) | prec {n f g} : @Primrec' n f → @Primrec' (n + 2) g → - Primrec' fun v : Vector ℕ (n + 1) => + Primrec' fun v : Mathlib.Vector ℕ (n + 1) => v.head.rec (f v.tail) fun y IH => g (y ::ᵥ IH ::ᵥ v.tail) end Nat @@ -1280,7 +1281,8 @@ theorem to_prim {n f} (pf : @Nat.Primrec' n f) : Primrec f := by Primrec.vector_cons.comp (Primrec.snd.comp .snd) <| (@Primrec.vector_tail _ _ (n + 1)).comp .fst).to₂ -theorem of_eq {n} {f g : Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : Primrec' g := +theorem of_eq {n} {f g : Mathlib.Vector ℕ n → ℕ} (hf : Primrec' f) (H : ∀ i, f i = g i) : + Primrec' g := (funext H : f = g) ▸ hf theorem const {n} : ∀ m, @Primrec' n fun _ => m @@ -1295,7 +1297,7 @@ theorem tail {n f} (hf : @Primrec' n f) : @Primrec' n.succ fun v => f v.tail := rw [← ofFn_get v.tail]; congr; funext i; simp /-- A function from vectors to vectors is primitive recursive when all of its projections are. -/ -def Vec {n m} (f : Vector ℕ n → Vector ℕ m) : Prop := +def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) : Prop := ∀ i, Primrec' fun v => (f v).get i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 @@ -1390,7 +1392,7 @@ theorem unpair₂ {n f} (hf : @Primrec' n f) : @Primrec' n fun v => (f v).unpair theorem of_prim {n f} : Primrec f → @Primrec' n f := suffices ∀ f, Nat.Primrec f → @Primrec' 1 fun v => f v.head from fun hf => (pred.comp₁ _ <| - (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (Vector ℕ n) _ m).map f) + (this _ hf).comp₁ (fun m => Encodable.encode <| (@decode (Mathlib.Vector ℕ n) _ m).map f) Primrec'.encode).of_eq fun i => by simp [encodek] fun f hf => by diff --git a/Mathlib/Computability/TuringMachine.lean b/Mathlib/Computability/TuringMachine.lean index 4e3298c29a939..378d518d12b0c 100644 --- a/Mathlib/Computability/TuringMachine.lean +++ b/Mathlib/Computability/TuringMachine.lean @@ -1443,7 +1443,7 @@ open TM1 variable {Γ : Type*} theorem exists_enc_dec [Inhabited Γ] [Finite Γ] : - ∃ (n : ℕ) (enc : Γ → Vector Bool n) (dec : Vector Bool n → Γ), + ∃ (n : ℕ) (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ), enc default = Vector.replicate n false ∧ ∀ a, dec (enc a) = a := by rcases Finite.exists_equiv_fin Γ with ⟨n, ⟨e⟩⟩ letI : DecidableEq Γ := e.decidableEq @@ -1475,13 +1475,13 @@ local notation "Stmt'₁" => Stmt Bool Λ'₁ σ local notation "Cfg'₁" => Cfg Bool Λ'₁ σ /-- Read a vector of length `n` from the tape. -/ -def readAux : ∀ n, (Vector Bool n → Stmt'₁) → Stmt'₁ +def readAux : ∀ n, (Mathlib.Vector Bool n → Stmt'₁) → Stmt'₁ | 0, f => f Vector.nil | i + 1, f => Stmt.branch (fun a _ ↦ a) (Stmt.move Dir.right <| readAux i fun v ↦ f (true ::ᵥ v)) (Stmt.move Dir.right <| readAux i fun v ↦ f (false ::ᵥ v)) -variable {n : ℕ} (enc : Γ → Vector Bool n) (dec : Vector Bool n → Γ) +variable {n : ℕ} (enc : Γ → Mathlib.Vector Bool n) (dec : Mathlib.Vector Bool n → Γ) /-- A move left or right corresponds to `n` moves across the super-cell. -/ def move (d : Dir) (q : Stmt'₁) : Stmt'₁ := @@ -1530,7 +1530,8 @@ theorem supportsStmt_write {S : Finset Λ'₁} {l : List Bool} {q : Stmt'₁} : theorem supportsStmt_read {S : Finset Λ'₁} : ∀ {f : Γ → Stmt'₁}, (∀ a, SupportsStmt S (f a)) → SupportsStmt S (read dec f) := suffices - ∀ (i) (f : Vector Bool i → Stmt'₁), (∀ v, SupportsStmt S (f v)) → SupportsStmt S (readAux i f) + ∀ (i) (f : Mathlib.Vector Bool i → Stmt'₁), + (∀ v, SupportsStmt S (f v)) → SupportsStmt S (readAux i f) from fun hf ↦ this n _ (by intro; simp only [supportsStmt_move, hf]) fun i f hf ↦ by induction' i with i IH; · exact hf _ diff --git a/Mathlib/Data/Fintype/BigOperators.lean b/Mathlib/Data/Fintype/BigOperators.lean index 6c8241d4dd465..0825cf4335415 100644 --- a/Mathlib/Data/Fintype/BigOperators.lean +++ b/Mathlib/Data/Fintype/BigOperators.lean @@ -172,7 +172,8 @@ theorem Fintype.card_fun [DecidableEq α] [Fintype α] [Fintype β] : simp @[simp] -theorem card_vector [Fintype α] (n : ℕ) : Fintype.card (Vector α n) = Fintype.card α ^ n := by +theorem card_vector [Fintype α] (n : ℕ) : + Fintype.card (Mathlib.Vector α n) = Fintype.card α ^ n := by rw [Fintype.ofEquiv_card]; simp /-- It is equivalent to compute the product of a function over `Fin n` or `Finset.range n`. -/ diff --git a/Mathlib/Data/Fintype/Fin.lean b/Mathlib/Data/Fintype/Fin.lean index 8744f61e2f3ae..96b7ea32bd858 100644 --- a/Mathlib/Data/Fintype/Fin.lean +++ b/Mathlib/Data/Fintype/Fin.lean @@ -61,7 +61,7 @@ theorem card_filter_univ_succ' (p : Fin (n + 1) → Prop) [DecidablePred p] : #{x | p x} = ite (p 0) 1 0 + #{x | p (.succ x)}:= by rw [card_filter_univ_succ]; split_ifs <;> simp [add_comm] -theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Vector α n) : +theorem card_filter_univ_eq_vector_get_eq_count [DecidableEq α] (a : α) (v : Mathlib.Vector α n) : #{i | v.get i = a} = v.toList.count a := by induction' v with n x xs hxs · simp diff --git a/Mathlib/Data/Fintype/Vector.lean b/Mathlib/Data/Fintype/Vector.lean index 2e94be07c8dbd..9b833c039d387 100644 --- a/Mathlib/Data/Fintype/Vector.lean +++ b/Mathlib/Data/Fintype/Vector.lean @@ -14,7 +14,7 @@ open Mathlib (Vector) variable {α : Type*} -instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (Vector α n) := +instance Vector.fintype [Fintype α] {n : ℕ} : Fintype (Mathlib.Vector α n) := Fintype.ofEquiv _ (Equiv.vectorEquivFin _ _).symm instance [DecidableEq α] [Fintype α] {n : ℕ} : Fintype (Sym.Sym' α n) := by diff --git a/Mathlib/Data/Num/Bitwise.lean b/Mathlib/Data/Num/Bitwise.lean index b374f02e13946..57703e62a4600 100644 --- a/Mathlib/Data/Num/Bitwise.lean +++ b/Mathlib/Data/Num/Bitwise.lean @@ -409,7 +409,7 @@ end SNum namespace SNum /-- `a.bits n` is the vector of the `n` first bits of `a` (starting from the LSB). -/ -def bits : SNum → ∀ n, Vector Bool n +def bits : SNum → ∀ n, Mathlib.Vector Bool n | _, 0 => Vector.nil | p, n + 1 => head p ::ᵥ bits (tail p) n diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index ac156160cdcb8..32aca18c1b5c4 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -55,7 +55,7 @@ instance {α : Type*} {n : ℕ} [DecidableEq α] : DecidableEq (Sym α n) := See note [reducible non-instances]. -/ -abbrev Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := +abbrev Mathlib.Vector.Perm.isSetoid (α : Type*) (n : ℕ) : Setoid (Vector α n) := (List.isSetoid α).comap Subtype.val attribute [local instance] Vector.Perm.isSetoid @@ -119,20 +119,21 @@ theorem coe_cons (s : Sym α n) (a : α) : (a ::ₛ s : Multiset α) = a ::ₘ s /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -def ofVector : Vector α n → Sym α n := +def ofVector : Mathlib.Vector α n → Sym α n := fun x => ⟨↑x.val, (Multiset.coe_card _).trans x.2⟩ /-- This is the quotient map that takes a list of n elements as an n-tuple and produces an nth symmetric power. -/ -instance : Coe (Vector α n) (Sym α n) where coe x := ofVector x +instance : Coe (Mathlib.Vector α n) (Sym α n) where coe x := ofVector x @[simp] -theorem ofVector_nil : ↑(Vector.nil : Vector α 0) = (Sym.nil : Sym α 0) := +theorem ofVector_nil : ↑(Vector.nil : Mathlib.Vector α 0) = (Sym.nil : Sym α 0) := rfl @[simp] -theorem ofVector_cons (a : α) (v : Vector α n) : ↑(Vector.cons a v) = a ::ₛ (↑v : Sym α n) := by +theorem ofVector_cons (a : α) (v : Mathlib.Vector α n) : + ↑(Vector.cons a v) = a ::ₛ (↑v : Sym α n) := by cases v rfl @@ -179,13 +180,13 @@ theorem mem_cons_of_mem (h : a ∈ s) : a ∈ b ::ₛ s := theorem mem_cons_self (a : α) (s : Sym α n) : a ∈ a ::ₛ s := Multiset.mem_cons_self a s.1 -theorem cons_of_coe_eq (a : α) (v : Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := +theorem cons_of_coe_eq (a : α) (v : Mathlib.Vector α n) : a ::ₛ (↑v : Sym α n) = ↑(a ::ᵥ v) := Subtype.ext <| by cases v rfl open scoped List in -theorem sound {a b : Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := +theorem sound {a b : Mathlib.Vector α n} (h : a.val ~ b.val) : (↑a : Sym α n) = ↑b := Subtype.ext <| Quotient.sound h /-- `erase s a h` is the sym that subtracts 1 from the diff --git a/Mathlib/Data/Sym/Sym2.lean b/Mathlib/Data/Sym/Sym2.lean index 28fbe8b4427a2..a827017219d56 100644 --- a/Mathlib/Data/Sym/Sym2.lean +++ b/Mathlib/Data/Sym/Sym2.lean @@ -586,7 +586,7 @@ section SymEquiv attribute [local instance] Vector.Perm.isSetoid -private def fromVector : Vector α 2 → α × α +private def fromVector : Mathlib.Vector α 2 → α × α | ⟨[a, b], _⟩ => (a, b) private theorem perm_card_two_iff {a₁ b₁ a₂ b₂ : α} : diff --git a/Mathlib/Logic/Equiv/List.lean b/Mathlib/Logic/Equiv/List.lean index c0753898bf0dc..6a6bb07e4ed21 100644 --- a/Mathlib/Logic/Equiv/List.lean +++ b/Mathlib/Logic/Equiv/List.lean @@ -127,11 +127,11 @@ noncomputable def _root_.Fintype.toEncodable (α : Type*) [Fintype α] : Encodab classical exact (Fintype.truncEncodable α).out /-- If `α` is encodable, then so is `Vector α n`. -/ -instance _root_.Vector.encodable [Encodable α] {n} : Encodable (Vector α n) := +instance Mathlib.Vector.encodable [Encodable α] {n} : Encodable (Mathlib.Vector α n) := Subtype.encodable /-- If `α` is countable, then so is `Vector α n`. -/ -instance _root_.Vector.countable [Countable α] {n} : Countable (Vector α n) := +instance Mathlib.Vector.countable [Countable α] {n} : Countable (Mathlib.Vector α n) := Subtype.countable /-- If `α` is encodable, then so is `Fin n → α`. -/ diff --git a/Mathlib/Logic/Small/List.lean b/Mathlib/Logic/Small/List.lean index 13b6fd0b0469a..f0706af610d10 100644 --- a/Mathlib/Logic/Small/List.lean +++ b/Mathlib/Logic/Small/List.lean @@ -18,9 +18,9 @@ universe u v open Mathlib -instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (Vector α n) := +instance smallVector {α : Type v} {n : ℕ} [Small.{u} α] : Small.{u} (Mathlib.Vector α n) := small_of_injective (Equiv.vectorEquivFin α n).injective instance smallList {α : Type v} [Small.{u} α] : Small.{u} (List α) := by - let e : (Σn, Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length + let e : (Σn, Mathlib.Vector α n) ≃ List α := Equiv.sigmaFiberEquiv List.length exact small_of_surjective e.surjective diff --git a/MathlibTest/fractran.lean b/MathlibTest/fractran.lean new file mode 100644 index 0000000000000..6df5e0aa16a87 --- /dev/null +++ b/MathlibTest/fractran.lean @@ -0,0 +1,54 @@ +import Mathlib.Data.PNat.Defs +import Mathlib.Tactic.NormNum + +open Nat + +def FRACTRAN := List (PNat × PNat) + +namespace FRACTRAN + +def mul? (n a b : PNat) : Option PNat := + if h : (n * a % b : Nat) = 0 then + some ⟨n * a / b, Nat.div_pos (le_of_dvd (Nat.mul_pos n.2 a.2) (dvd_of_mod_eq_zero h)) b.2⟩ + else + none + +def step (f : FRACTRAN) (n : PNat) : Option PNat := + f.findSome? fun ⟨a, b⟩ => mul? n a b + +def listStep (f : FRACTRAN) (n : List PNat) : List PNat := + match n with + | [] => [] + | n :: ns => + match step f n with + | some n' => n' :: n :: ns + | none => n :: ns + +def runSteps (f : FRACTRAN) (n : PNat) (k : Nat) : List PNat := + List.range k |>.foldl (fun r _ => listStep f r) [n] + +def runFor (f : FRACTRAN) (n : PNat) : Nat → PNat + | 0 => n + | k + 1 => match step f n with + | some n' => runFor f n' k + | none => n + +-- Check that `[(3, 2)]` is the program for addition. +example : (runFor [(3, 2)] ⟨2^37 * 3^42, by norm_num⟩ 1000).val = 3^(37+42) := rfl + +def PRIMEGAME : FRACTRAN := [(17, 91), (78, 85), (19, 51), (23, 38), (29, 33), (77, 29), (95, 23), + (77, 19), (1, 17), (11, 13), (13, 11), (15, 2), (1, 7), (55, 1)] + +def twoPow? (n : Nat) : Option Nat := + if _ : n = 0 then none else + if _ : n = 1 then some 0 else + if h : n % 2 = 0 then + twoPow? (n / 2) |>.map (· + 1) + else + none + +example : twoPow? 1024 = some 10 := by native_decide + +#eval runSteps PRIMEGAME ⟨2, by norm_num⟩ 100000 |>.map PNat.val|>.filterMap twoPow? + +end FRACTRAN From 5470a0ca6767e9aeb61cef8d5ab0cce11f4197c5 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 21:50:08 +1100 Subject: [PATCH 132/250] disambiguate Vector --- Mathlib/Computability/Halting.lean | 28 +++++++++++++----------- Mathlib/Computability/Partrec.lean | 19 ++++++++-------- Mathlib/Computability/TMToPartrec.lean | 16 ++++++++------ Mathlib/GroupTheory/Perm/Cycle/Type.lean | 10 ++++----- Mathlib/SetTheory/Cardinal/Basic.lean | 4 ++-- Mathlib/Topology/List.lean | 20 ++++++++--------- 6 files changed, 51 insertions(+), 46 deletions(-) diff --git a/Mathlib/Computability/Halting.lean b/Mathlib/Computability/Halting.lean index 0bce012c8ecc0..7437dbc1b31ac 100644 --- a/Mathlib/Computability/Halting.lean +++ b/Mathlib/Computability/Halting.lean @@ -267,11 +267,12 @@ namespace Nat open Vector Part /-- A simplified basis for `Partrec`. -/ -inductive Partrec' : ∀ {n}, (Vector ℕ n →. ℕ) → Prop +inductive Partrec' : ∀ {n}, (Mathlib.Vector ℕ n →. ℕ) → Prop | prim {n f} : @Primrec' n f → @Partrec' n f - | comp {m n f} (g : Fin n → Vector ℕ m →. ℕ) : - Partrec' f → (∀ i, Partrec' (g i)) → Partrec' fun v => (Vector.mOfFn fun i => g i v) >>= f - | rfind {n} {f : Vector ℕ (n + 1) → ℕ} : + | comp {m n f} (g : Fin n → Mathlib.Vector ℕ m →. ℕ) : + Partrec' f → (∀ i, Partrec' (g i)) → + Partrec' fun v => (Mathlib.Vector.mOfFn fun i => g i v) >>= f + | rfind {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} : @Partrec' (n + 1) f → Partrec' fun v => rfind fun n => some (f (n ::ᵥ v) = 0) end Nat @@ -293,10 +294,11 @@ theorem to_part {n f} (pf : @Partrec' n f) : _root_.Partrec f := by this).to₂.partrec₂ exact _root_.Partrec.rfind this -theorem of_eq {n} {f g : Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : Partrec' g := +theorem of_eq {n} {f g : Mathlib.Vector ℕ n →. ℕ} (hf : Partrec' f) (H : ∀ i, f i = g i) : + Partrec' g := (funext H : f = g) ▸ hf -theorem of_prim {n} {f : Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := +theorem of_prim {n} {f : Mathlib.Vector ℕ n → ℕ} (hf : Primrec f) : @Partrec' n f := prim (Nat.Primrec'.of_prim hf) theorem head {n : ℕ} : @Partrec' n.succ (@head ℕ n) := @@ -313,21 +315,21 @@ protected theorem bind {n f g} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : exact prim (Nat.Primrec'.get _)).of_eq fun v => by simp [mOfFn, Part.bind_assoc, pure] -protected theorem map {n f} {g : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) +protected theorem map {n f} {g : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' n f) (hg : @Partrec' (n + 1) g) : @Partrec' n fun v => (f v).map fun a => g (a ::ᵥ v) := by simpa [(Part.bind_some_eq_map _ _).symm] using hf.bind hg /-- Analogous to `Nat.Partrec'` for `ℕ`-valued functions, a predicate for partial recursive vector-valued functions. -/ -def Vec {n m} (f : Vector ℕ n → Vector ℕ m) := +def Vec {n m} (f : Mathlib.Vector ℕ n → Mathlib.Vector ℕ m) := ∀ i, Partrec' fun v => (f v).get i nonrec theorem Vec.prim {n m f} (hf : @Nat.Primrec'.Vec n m f) : Vec f := fun i => prim <| hf i protected theorem nil {n} : @Vec n 0 fun _ => nil := fun i => i.elim0 -protected theorem cons {n m} {f : Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) (hg : @Vec n m g) : - Vec fun v => f v ::ᵥ g v := fun i => +protected theorem cons {n m} {f : Mathlib.Vector ℕ n → ℕ} {g} (hf : @Partrec' n f) + (hg : @Vec n m g) : Vec fun v => f v ::ᵥ g v := fun i => Fin.cases (by simpa using hf) (fun i => by simp only [hg i, get_cons_succ]) i theorem idv {n} : @Vec n n id := @@ -336,11 +338,11 @@ theorem idv {n} : @Vec n n id := theorem comp' {n m f g} (hf : @Partrec' m f) (hg : @Vec n m g) : Partrec' fun v => f (g v) := (hf.comp _ hg).of_eq fun v => by simp -theorem comp₁ {n} (f : ℕ →. ℕ) {g : Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) +theorem comp₁ {n} (f : ℕ →. ℕ) {g : Mathlib.Vector ℕ n → ℕ} (hf : @Partrec' 1 fun v => f v.head) (hg : @Partrec' n g) : @Partrec' n fun v => f (g v) := by simpa using hf.comp' (Partrec'.cons hg Partrec'.nil) -theorem rfindOpt {n} {f : Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : +theorem rfindOpt {n} {f : Mathlib.Vector ℕ (n + 1) → ℕ} (hf : @Partrec' (n + 1) f) : @Partrec' n fun v => Nat.rfindOpt fun a => ofNat (Option ℕ) (f (a ::ᵥ v)) := ((rfind <| (of_prim (Primrec.nat_sub.comp (_root_.Primrec.const 1) Primrec.vector_head)).comp₁ @@ -367,7 +369,7 @@ open Nat.Partrec.Code theorem of_part : ∀ {n f}, _root_.Partrec f → @Partrec' n f := @(suffices ∀ f, Nat.Partrec f → @Partrec' 1 fun v => f v.head from fun {n f} hf => by let g := fun n₁ => - (Part.ofOption (decode (α := Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) + (Part.ofOption (decode (α := Mathlib.Vector ℕ n) n₁)).bind (fun a => Part.map encode (f a)) exact (comp₁ g (this g hf) (prim Nat.Primrec'.encode)).of_eq fun i => by dsimp only [g]; simp [encodek, Part.map_id'] diff --git a/Mathlib/Computability/Partrec.lean b/Mathlib/Computability/Partrec.lean index 63fdfb8565e1c..3ad33d20e77c5 100644 --- a/Mathlib/Computability/Partrec.lean +++ b/Mathlib/Computability/Partrec.lean @@ -314,25 +314,25 @@ theorem list_concat : Computable₂ fun l (a : α) => l ++ [a] := theorem list_length : Computable (@List.length α) := Primrec.list_length.to_comp -theorem vector_cons {n} : Computable₂ (@Vector.cons α n) := +theorem vector_cons {n} : Computable₂ (@Mathlib.Vector.cons α n) := Primrec.vector_cons.to_comp -theorem vector_toList {n} : Computable (@Vector.toList α n) := +theorem vector_toList {n} : Computable (@Mathlib.Vector.toList α n) := Primrec.vector_toList.to_comp -theorem vector_length {n} : Computable (@Vector.length α n) := +theorem vector_length {n} : Computable (@Mathlib.Vector.length α n) := Primrec.vector_length.to_comp -theorem vector_head {n} : Computable (@Vector.head α n) := +theorem vector_head {n} : Computable (@Mathlib.Vector.head α n) := Primrec.vector_head.to_comp -theorem vector_tail {n} : Computable (@Vector.tail α n) := +theorem vector_tail {n} : Computable (@Mathlib.Vector.tail α n) := Primrec.vector_tail.to_comp -theorem vector_get {n} : Computable₂ (@Vector.get α n) := +theorem vector_get {n} : Computable₂ (@Mathlib.Vector.get α n) := Primrec.vector_get.to_comp -theorem vector_ofFn' {n} : Computable (@Vector.ofFn α n) := +theorem vector_ofFn' {n} : Computable (@Mathlib.Vector.ofFn α n) := Primrec.vector_ofFn'.to_comp theorem fin_app {n} : Computable₂ (@id (Fin n → σ)) := @@ -523,7 +523,8 @@ end Partrec @[simp] theorem Vector.mOfFn_part_some {α n} : - ∀ f : Fin n → α, (Vector.mOfFn fun i => Part.some (f i)) = Part.some (Vector.ofFn f) := + ∀ f : Fin n → α, + (Mathlib.Vector.mOfFn fun i => Part.some (f i)) = Part.some (Mathlib.Vector.ofFn f) := Vector.mOfFn_pure namespace Computable @@ -642,7 +643,7 @@ theorem list_ofFn : exact list_cons.comp (hf 0) (list_ofFn fun i => hf i.succ) theorem vector_ofFn {n} {f : Fin n → α → σ} (hf : ∀ i, Computable (f i)) : - Computable fun a => Vector.ofFn fun i => f i a := + Computable fun a => Mathlib.Vector.ofFn fun i => f i a := (Partrec.vector_mOfFn hf).of_eq fun a => by simp end Computable diff --git a/Mathlib/Computability/TMToPartrec.lean b/Mathlib/Computability/TMToPartrec.lean index 17abd70e3ad66..effb93335081f 100644 --- a/Mathlib/Computability/TMToPartrec.lean +++ b/Mathlib/Computability/TMToPartrec.lean @@ -239,12 +239,14 @@ def prec (f g : Code) : Code := attribute [-simp] Part.bind_eq_bind Part.map_eq_map Part.pure_eq_some -theorem exists_code.comp {m n} {f : Vector ℕ n →. ℕ} {g : Fin n → Vector ℕ m →. ℕ} - (hf : ∃ c : Code, ∀ v : Vector ℕ n, c.eval v.1 = pure <$> f v) - (hg : ∀ i, ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = pure <$> g i v) : - ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = pure <$> ((Vector.mOfFn fun i => g i v) >>= f) := by +theorem exists_code.comp {m n} {f : Mathlib.Vector ℕ n →. ℕ} {g : Fin n → Mathlib.Vector ℕ m →. ℕ} + (hf : ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v) + (hg : ∀ i, ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, c.eval v.1 = pure <$> g i v) : + ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, + c.eval v.1 = pure <$> ((Mathlib.Vector.mOfFn fun i => g i v) >>= f) := by rsuffices ⟨cg, hg⟩ : - ∃ c : Code, ∀ v : Vector ℕ m, c.eval v.1 = Subtype.val <$> Vector.mOfFn fun i => g i v + ∃ c : Code, ∀ v : Mathlib.Vector ℕ m, + c.eval v.1 = Subtype.val <$> Mathlib.Vector.mOfFn fun i => g i v · obtain ⟨cf, hf⟩ := hf exact ⟨cf.comp cg, fun v => by @@ -259,8 +261,8 @@ theorem exists_code.comp {m n} {f : Vector ℕ n →. ℕ} {g : Fin n → Vector simp [Vector.mOfFn, hg₁, map_bind, seq_bind_eq, bind_assoc, (· ∘ ·), hl] rfl⟩ -theorem exists_code {n} {f : Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : - ∃ c : Code, ∀ v : Vector ℕ n, c.eval v.1 = pure <$> f v := by +theorem exists_code {n} {f : Mathlib.Vector ℕ n →. ℕ} (hf : Nat.Partrec' f) : + ∃ c : Code, ∀ v : Mathlib.Vector ℕ n, c.eval v.1 = pure <$> f v := by induction hf with | prim hf => induction hf with diff --git a/Mathlib/GroupTheory/Perm/Cycle/Type.lean b/Mathlib/GroupTheory/Perm/Cycle/Type.lean index 7b8e2f1c0b26d..b47673b955137 100644 --- a/Mathlib/GroupTheory/Perm/Cycle/Type.lean +++ b/Mathlib/GroupTheory/Perm/Cycle/Type.lean @@ -394,12 +394,12 @@ section Cauchy variable (G : Type*) [Group G] (n : ℕ) /-- The type of vectors with terms from `G`, length `n`, and product equal to `1:G`. -/ -def vectorsProdEqOne : Set (Vector G n) := +def vectorsProdEqOne : Set (Mathlib.Vector G n) := { v | v.toList.prod = 1 } namespace VectorsProdEqOne -theorem mem_iff {n : ℕ} (v : Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := +theorem mem_iff {n : ℕ} (v : Mathlib.Vector G n) : v ∈ vectorsProdEqOne G n ↔ v.toList.prod = 1 := Iff.rfl theorem zero_eq : vectorsProdEqOne G 0 = {Vector.nil} := @@ -421,7 +421,7 @@ instance oneUnique : Unique (vectorsProdEqOne G 1) := by /-- Given a vector `v` of length `n`, make a vector of length `n + 1` whose product is `1`, by appending the inverse of the product of `v`. -/ @[simps] -def vectorEquiv : Vector G n ≃ vectorsProdEqOne G (n + 1) where +def vectorEquiv : Mathlib.Vector G n ≃ vectorsProdEqOne G (n + 1) where toFun v := ⟨v.toList.prod⁻¹ ::ᵥ v, by rw [mem_iff, Vector.toList_cons, List.prod_cons, inv_mul_cancel]⟩ invFun v := v.1.tail @@ -436,12 +436,12 @@ def vectorEquiv : Vector G n ≃ vectorsProdEqOne G (n + 1) where /-- Given a vector `v` of length `n` whose product is 1, make a vector of length `n - 1`, by deleting the last entry of `v`. -/ -def equivVector : ∀ n, vectorsProdEqOne G n ≃ Vector G (n - 1) +def equivVector : ∀ n, vectorsProdEqOne G n ≃ Mathlib.Vector G (n - 1) | 0 => (equivOfUnique (vectorsProdEqOne G 0) (vectorsProdEqOne G 1)).trans (vectorEquiv G 0).symm | (n + 1) => (vectorEquiv G n).symm instance [Fintype G] : Fintype (vectorsProdEqOne G n) := - Fintype.ofEquiv (Vector G (n - 1)) (equivVector G n).symm + Fintype.ofEquiv (Mathlib.Vector G (n - 1)) (equivVector G n).symm theorem card [Fintype G] : Fintype.card (vectorsProdEqOne G n) = Fintype.card G ^ (n - 1) := (Fintype.card_congr (equivVector G n)).trans (card_vector (n - 1)) diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 2b76c69c48ddc..ddae9a4470804 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -1779,12 +1779,12 @@ theorem mk_plift_false : #(PLift False) = 0 := mk_eq_zero _ @[simp] -theorem mk_vector (α : Type u) (n : ℕ) : #(Vector α n) = #α ^ n := +theorem mk_vector (α : Type u) (n : ℕ) : #(Mathlib.Vector α n) = #α ^ n := (mk_congr (Equiv.vectorEquivFin α n)).trans <| by simp theorem mk_list_eq_sum_pow (α : Type u) : #(List α) = sum fun n : ℕ => #α ^ n := calc - #(List α) = #(Σn, Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm + #(List α) = #(Σn, Mathlib.Vector α n) := mk_congr (Equiv.sigmaFiberEquiv List.length).symm _ = sum fun n : ℕ => #α ^ n := by simp theorem mk_quot_le {α : Type u} {r : α → α → Prop} : #(Quot r) ≤ #α := diff --git a/Mathlib/Topology/List.lean b/Mathlib/Topology/List.lean index 0620deab85315..4ca0d9581282a 100644 --- a/Mathlib/Topology/List.lean +++ b/Mathlib/Topology/List.lean @@ -184,16 +184,16 @@ open List open Mathlib (Vector) -instance (n : ℕ) : TopologicalSpace (Vector α n) := by unfold Vector; infer_instance +instance (n : ℕ) : TopologicalSpace (Mathlib.Vector α n) := by unfold Mathlib.Vector; infer_instance -theorem tendsto_cons {n : ℕ} {a : α} {l : Vector α n} : - Tendsto (fun p : α × Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by +theorem tendsto_cons {n : ℕ} {a : α} {l : Mathlib.Vector α n} : + Tendsto (fun p : α × Mathlib.Vector α n => p.1 ::ᵥ p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (a ::ᵥ l)) := by rw [tendsto_subtype_rng, Vector.cons_val] exact tendsto_fst.cons (Tendsto.comp continuousAt_subtype_val tendsto_snd) theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : - ∀ {l : Vector α n}, - Tendsto (fun p : α × Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) + ∀ {l : Mathlib.Vector α n}, + Tendsto (fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2) (𝓝 a ×ˢ 𝓝 l) (𝓝 (Vector.insertIdx a i l)) | ⟨l, hl⟩ => by rw [Vector.insertIdx, tendsto_subtype_rng] @@ -204,29 +204,29 @@ theorem tendsto_insertIdx {n : ℕ} {i : Fin (n + 1)} {a : α} : /-- Continuity of `Vector.insertIdx`. -/ theorem continuous_insertIdx' {n : ℕ} {i : Fin (n + 1)} : - Continuous fun p : α × Vector α n => Vector.insertIdx p.1 i p.2 := + Continuous fun p : α × Mathlib.Vector α n => Vector.insertIdx p.1 i p.2 := continuous_iff_continuousAt.mpr fun ⟨a, l⟩ => by rw [ContinuousAt, nhds_prod_eq]; exact tendsto_insertIdx @[deprecated (since := "2024-10-21")] alias continuous_insertNth' := continuous_insertIdx' -theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Vector α n} +theorem continuous_insertIdx {n : ℕ} {i : Fin (n + 1)} {f : β → α} {g : β → Mathlib.Vector α n} (hf : Continuous f) (hg : Continuous g) : Continuous fun b => Vector.insertIdx (f b) i (g b) := continuous_insertIdx'.comp (hf.prod_mk hg : _) @[deprecated (since := "2024-10-21")] alias continuous_insertNth := continuous_insertIdx theorem continuousAt_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - ∀ {l : Vector α (n + 1)}, ContinuousAt (Vector.eraseIdx i) l + ∀ {l : Mathlib.Vector α (n + 1)}, ContinuousAt (Mathlib.Vector.eraseIdx i) l | ⟨l, hl⟩ => by - rw [ContinuousAt, Vector.eraseIdx, tendsto_subtype_rng] + rw [ContinuousAt, Mathlib.Vector.eraseIdx, tendsto_subtype_rng] simp only [Vector.eraseIdx_val] exact Tendsto.comp List.tendsto_eraseIdx continuousAt_subtype_val @[deprecated (since := "2024-05-04")] alias continuousAt_removeNth := continuousAt_eraseIdx theorem continuous_eraseIdx {n : ℕ} {i : Fin (n + 1)} : - Continuous (Vector.eraseIdx i : Vector α (n + 1) → Vector α n) := + Continuous (Mathlib.Vector.eraseIdx i : Mathlib.Vector α (n + 1) → Mathlib.Vector α n) := continuous_iff_continuousAt.mpr fun ⟨_a, _l⟩ => continuousAt_eraseIdx @[deprecated (since := "2024-05-04")] alias continuous_removeNth := continuous_eraseIdx From a151e519ad7822530ddb3fb25275d41aa3a80c55 Mon Sep 17 00:00:00 2001 From: su00000 Date: Mon, 25 Nov 2024 11:04:38 +0000 Subject: [PATCH 133/250] feat(RingTheory/LocalProperties): Add theorems about LocalizedModule (#19080) This PR contains: 1. A proof that localization commutes with ranges. 2. Some missing lemmas about localized submodules. 3. A proof that injective, surjective, bijective (of linear maps) are local properties. These theorems will be used to prove that flat is a local property in a later PR. Co-authored-by: Yongle Hu Co-authored-by: Yi Song Co-authored-by: su00000 --- Mathlib.lean | 1 + .../Module/LocalizedModule/Submodule.lean | 56 ++--- Mathlib/RingTheory/Ideal/Maximal.lean | 9 + .../RingTheory/LocalProperties/Exactness.lean | 194 ++++++++++++++++++ .../RingTheory/LocalProperties/Submodule.lean | 81 +++++++- 5 files changed, 316 insertions(+), 25 deletions(-) create mode 100644 Mathlib/RingTheory/LocalProperties/Exactness.lean diff --git a/Mathlib.lean b/Mathlib.lean index 4b515136d2aa3..2c97f33123aff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4322,6 +4322,7 @@ import Mathlib.RingTheory.LaurentSeries import Mathlib.RingTheory.LinearDisjoint import Mathlib.RingTheory.LittleWedderburn import Mathlib.RingTheory.LocalProperties.Basic +import Mathlib.RingTheory.LocalProperties.Exactness import Mathlib.RingTheory.LocalProperties.IntegrallyClosed import Mathlib.RingTheory.LocalProperties.Projective import Mathlib.RingTheory.LocalProperties.Reduced diff --git a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean index 650844736f7ca..0ec7f25bf02d7 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/Submodule.lean @@ -189,31 +189,32 @@ instance (M' : Submodule R M) : IsLocalizedModule p (M'.toLocalizedQuotient p) : end Quotient -section LinearMap +namespace LinearMap -variable {P : Type*} [AddCommGroup P] [Module R P] -variable {Q : Type*} [AddCommGroup Q] [Module R Q] [Module S Q] [IsScalarTower R S Q] +variable {P : Type*} [AddCommMonoid P] [Module R P] +variable {Q : Type*} [AddCommMonoid 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) : - Submodule.localized' S p f (LinearMap.ker g) = - LinearMap.ker ((IsLocalizedModule.map p f f' g).extendScalarsOfIsLocalization p S) := by +open Submodule IsLocalizedModule + +lemma ker_localizedMap_eq_localized₀_ker (g : M →ₗ[R] P) : + ker (map p f f' g) = (ker g).localized₀ p f := by ext x - simp only [Submodule.mem_localized', mem_ker, extendScalarsOfIsLocalization_apply'] - constructor - · rintro ⟨m, hm, a, ha, rfl⟩ - rw [IsLocalizedModule.map_mk', hm] - simp - · intro h - obtain ⟨⟨a, b⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x - simp only [Function.uncurry_apply_pair, IsLocalizedModule.map_mk', - IsLocalizedModule.mk'_eq_zero, IsLocalizedModule.eq_zero_iff p f'] at h + simp only [Submodule.mem_localized₀, mem_ker] + refine ⟨fun h ↦ ?_, ?_⟩ + · obtain ⟨⟨a, b⟩, rfl⟩ := IsLocalizedModule.mk'_surjective p f x + simp only [Function.uncurry_apply_pair, map_mk', mk'_eq_zero, eq_zero_iff p f'] at h obtain ⟨c, hc⟩ := h refine ⟨c • a, by simpa, c * b, by simp⟩ + · rintro ⟨m, hm, a, ha, rfl⟩ + simp [IsLocalizedModule.map_mk', hm] -lemma LinearMap.ker_localizedMap_eq_localized'_ker (g : M →ₗ[R] P) : - LinearMap.ker (IsLocalizedModule.map p f f' g) = - ((LinearMap.ker g).localized' S p f).restrictScalars _ := by +lemma localized'_ker_eq_ker_localizedMap (g : M →ₗ[R] P) : + (ker g).localized' S p f = ker ((map p f f' g).extendScalarsOfIsLocalization p S) := + SetLike.ext (by apply SetLike.ext_iff.mp (f.ker_localizedMap_eq_localized₀_ker p f' g).symm) + +lemma ker_localizedMap_eq_localized'_ker (g : M →ₗ[R] P) : + ker (map p f f' g) = ((ker g).localized' S p f).restrictScalars _ := by ext simp [localized'_ker_eq_ker_localizedMap S p f f'] @@ -223,18 +224,27 @@ The canonical map from the kernel of `g` to the kernel of `g` localized at a sub This is a localization map by `LinearMap.toKerLocalized_isLocalizedModule`. -/ @[simps!] -noncomputable def LinearMap.toKerIsLocalized (g : M →ₗ[R] P) : - ker g →ₗ[R] ker (IsLocalizedModule.map p f f' g) := - f.restrict (fun x hx ↦ by simp [LinearMap.mem_ker, LinearMap.mem_ker.mp hx]) +noncomputable def toKerIsLocalized (g : M →ₗ[R] P) : + ker g →ₗ[R] ker (map p f f' g) := + f.restrict (fun x hx ↦ by simp [mem_ker, mem_ker.mp hx]) include S in /-- The canonical map to the kernel of the localization of `g` is localizing. In other words, localization commutes with kernels. -/ -lemma LinearMap.toKerLocalized_isLocalizedModule (g : M →ₗ[R] P) : +lemma toKerLocalized_isLocalizedModule (g : M →ₗ[R] P) : IsLocalizedModule p (toKerIsLocalized p f f' g) := let e : Submodule.localized' S p f (ker g) ≃ₗ[S] - ker ((IsLocalizedModule.map p f f' g).extendScalarsOfIsLocalization p S) := + ker ((map p f f' g).extendScalarsOfIsLocalization p S) := LinearEquiv.ofEq _ _ (localized'_ker_eq_ker_localizedMap S p f f' g) IsLocalizedModule.of_linearEquiv p (Submodule.toLocalized' S p f (ker g)) (e.restrictScalars R) +lemma range_localizedMap_eq_localized₀_range (g : M →ₗ[R] P) : + range (map p f f' g) = (range g).localized₀ p f' := by + ext; simp [mem_localized₀, mem_range, (mk'_surjective p f).exists] + +/-- Localization commutes with ranges. -/ +lemma localized'_range_eq_range_localizedMap (g : M →ₗ[R] P) : + (range g).localized' S p f' = range ((map p f f' g).extendScalarsOfIsLocalization p S) := + SetLike.ext (by apply SetLike.ext_iff.mp (f.range_localizedMap_eq_localized₀_range p f' g).symm) + end LinearMap diff --git a/Mathlib/RingTheory/Ideal/Maximal.lean b/Mathlib/RingTheory/Ideal/Maximal.lean index 591bb5e34a064..444898cdd4f69 100644 --- a/Mathlib/RingTheory/Ideal/Maximal.lean +++ b/Mathlib/RingTheory/Ideal/Maximal.lean @@ -162,6 +162,15 @@ theorem IsMaximal.isPrime {I : Ideal α} (H : I.IsMaximal) : I.IsPrime := instance (priority := 100) IsMaximal.isPrime' (I : Ideal α) : ∀ [_H : I.IsMaximal], I.IsPrime := @IsMaximal.isPrime _ _ _ +theorem exists_disjoint_powers_of_span_eq_top (s : Set α) (hs : span s = ⊤) (I : Ideal α) + (hI : I ≠ ⊤) : ∃ r ∈ s, Disjoint (I : Set α) (Submonoid.powers r) := by + have ⟨M, hM, le⟩ := exists_le_maximal I hI + have := hM.1.1 + rw [Ne, eq_top_iff, ← hs, span_le, Set.not_subset] at this + have ⟨a, has, haM⟩ := this + exact ⟨a, has, Set.disjoint_left.mpr fun x hx ⟨n, hn⟩ ↦ + haM (hM.isPrime.mem_of_pow_mem _ (le <| hn ▸ hx))⟩ + theorem span_singleton_lt_span_singleton [IsDomain α] {x y : α} : span ({x} : Set α) < span ({y} : Set α) ↔ DvdNotUnit y x := by rw [lt_iff_le_not_le, span_singleton_le_span_singleton, span_singleton_le_span_singleton, diff --git a/Mathlib/RingTheory/LocalProperties/Exactness.lean b/Mathlib/RingTheory/LocalProperties/Exactness.lean new file mode 100644 index 0000000000000..3c6ae6bb837c7 --- /dev/null +++ b/Mathlib/RingTheory/LocalProperties/Exactness.lean @@ -0,0 +1,194 @@ +/- +Copyright (c) 2024 Sihan Su. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sihan Su, Yongle Hu, Yi Song +-/ +import Mathlib.Algebra.Exact +import Mathlib.RingTheory.LocalProperties.Submodule +import Mathlib.RingTheory.Localization.Away.Basic + +/-! +# Local properties about linear maps + +In this file, we show that +injectivity, surjectivity, bijectivity and exactness of linear maps are local properties. +More precisely, we show that these can be checked at maximal ideals and on standard covers. + +-/ +open Submodule LocalizedModule Ideal LinearMap + +section isLocalized_maximal + +open IsLocalizedModule + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] + +variable + (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)] + (Nₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Nₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Nₚ P)] + (g : ∀ (P : Ideal R) [P.IsMaximal], N →ₗ[R] Nₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (g P)] + (Lₚ : ∀ (P : Ideal R) [P.IsMaximal], Type*) + [∀ (P : Ideal R) [P.IsMaximal], AddCommMonoid (Lₚ P)] + [∀ (P : Ideal R) [P.IsMaximal], Module R (Lₚ P)] + (h : ∀ (P : Ideal R) [P.IsMaximal], L →ₗ[R] Lₚ P) + [∀ (P : Ideal R) [P.IsMaximal], IsLocalizedModule P.primeCompl (h P)] + (F : M →ₗ[R] N) (G : N →ₗ[R] L) + +theorem injective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Injective (map P.primeCompl (f P) (g P) F)) : + Function.Injective F := + fun x y eq ↦ Module.eq_of_localization_maximal _ f _ _ fun P _ ↦ H P <| by simp [eq] + +theorem surjective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Surjective (map P.primeCompl (f P) (g P) F)) : + Function.Surjective F := + range_eq_top.mp <| eq_top_of_localization₀_maximal Nₚ g _ <| + fun P _ ↦ (range_localizedMap_eq_localized₀_range _ (f P) (g P) F).symm.trans <| + range_eq_top.mpr <| H P + +theorem bijective_of_isLocalized_maximal + (H : ∀ (P : Ideal R) [P.IsMaximal], Function.Bijective (map P.primeCompl (f P) (g P) F)) : + Function.Bijective F := + ⟨injective_of_isLocalized_maximal Mₚ f Nₚ g F fun J _ ↦ (H J).1, + surjective_of_isLocalized_maximal Mₚ f Nₚ g F fun J _ ↦ (H J).2⟩ + +theorem exact_of_isLocalized_maximal (H : ∀ (J : Ideal R) [J.IsMaximal], + Function.Exact (map J.primeCompl (f J) (g J) F) (map J.primeCompl (g J) (h J) G)) : + Function.Exact F G := by + simp only [LinearMap.exact_iff] at H ⊢ + apply eq_of_localization₀_maximal Nₚ g + intro J hJ + rw [← LinearMap.range_localizedMap_eq_localized₀_range _ (f J) (g J) F, + ← LinearMap.ker_localizedMap_eq_localized₀_ker J.primeCompl (g J) (h J) G] + have := SetLike.ext_iff.mp <| H J + ext x + simp only [mem_range, mem_ker] at this ⊢ + exact this x + +end isLocalized_maximal + +section localized_maximal + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] (f : M →ₗ[R] N) (g : N →ₗ[R] L) + +theorem injective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Injective (map J.primeCompl f)) : + Function.Injective f := + injective_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) f h + +theorem surjective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Surjective (map J.primeCompl f)) : + Function.Surjective f := + surjective_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) f h + +theorem bijective_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Bijective (map J.primeCompl f)) : + Function.Bijective f := + ⟨injective_of_localized_maximal _ fun J _ ↦ (h J).1, + surjective_of_localized_maximal _ fun J _ ↦ (h J).2⟩ + +theorem exact_of_localized_maximal + (h : ∀ (J : Ideal R) [J.IsMaximal], Function.Exact (map J.primeCompl f) (map J.primeCompl g)) : + Function.Exact f g := + exact_of_isLocalized_maximal _ (fun _ _ ↦ mkLinearMap _ _) _ (fun _ _ ↦ mkLinearMap _ _) + _ (fun _ _ ↦ mkLinearMap _ _) f g h + +end localized_maximal + +section isLocalized_span + +open IsLocalizedModule + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] (s : Set R) (spn : Ideal.span s = ⊤) +include spn + +variable + (Mₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Mₚ r)] + [∀ r : s, Module R (Mₚ r)] + (f : ∀ r : s, M →ₗ[R] Mₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (f r)] + (Nₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Nₚ r)] + [∀ r : s, Module R (Nₚ r)] + (g : ∀ r : s, N →ₗ[R] Nₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (g r)] + (Lₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Lₚ r)] + [∀ r : s, Module R (Lₚ r)] + (h : ∀ r : s, L →ₗ[R] Lₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (h r)] + (F : M →ₗ[R] N) (G : N →ₗ[R] L) + +theorem injective_of_isLocalized_span + (H : ∀ r : s, Function.Injective (map (.powers r.1) (f r) (g r) F)) : + Function.Injective F := + fun x y eq ↦ Module.eq_of_isLocalized_span _ spn _ f _ _ fun P ↦ H P <| by simp [eq] + +theorem surjective_of_isLocalized_span + (H : ∀ r : s, Function.Surjective (map (.powers r.1) (f r) (g r) F)) : + Function.Surjective F := + range_eq_top.mp <| eq_top_of_isLocalized₀_span s spn Nₚ g fun r ↦ + (range_localizedMap_eq_localized₀_range _ (f r) (g r) F).symm.trans <| range_eq_top.mpr <| H r + +theorem bijective_of_isLocalized_span + (H : ∀ r : s, Function.Bijective (map (.powers r.1) (f r) (g r) F)) : + Function.Bijective F := + ⟨injective_of_isLocalized_span _ spn Mₚ f Nₚ g F fun r ↦ (H r).1, + surjective_of_isLocalized_span _ spn Mₚ f Nₚ g F fun r ↦ (H r).2⟩ + +lemma exact_of_isLocalized_span (H : ∀ r : s, Function.Exact + (map (.powers r.1) (f r) (g r) F) (map (.powers r.1) (g r) (h r) G)) : + Function.Exact F G := by + simp only [LinearMap.exact_iff] at H ⊢ + apply Submodule.eq_of_isLocalized₀_span s spn Nₚ g + intro r + rw [← LinearMap.range_localizedMap_eq_localized₀_range _ (f r) (g r) F] + rw [← LinearMap.ker_localizedMap_eq_localized₀_ker (.powers r.1) (g r) (h r) G] + have := SetLike.ext_iff.mp <| H r + ext x + simp only [mem_range, mem_ker] at this ⊢ + exact this x + +end isLocalized_span + +section localized_span + +variable {R M N L : Type*} [CommSemiring R] [AddCommMonoid M] [Module R M] + [AddCommMonoid N] [Module R N] [AddCommMonoid L] [Module R L] + (s : Set R) (spn : span s = ⊤) (f : M →ₗ[R] N) (g : N →ₗ[R] L) +include spn + +theorem injective_of_localized_span + (h : ∀ r : s, Function.Injective (map (.powers r.1) f)) : + Function.Injective f := + injective_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) f h + +theorem surjective_of_localized_span + (h : ∀ r : s, Function.Surjective (map (.powers r.1) f)) : + Function.Surjective f := + surjective_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) f h + +theorem bijective_of_localized_span + (h : ∀ r : s, Function.Bijective (map (.powers r.1) f)) : + Function.Bijective f := + ⟨injective_of_localized_span _ spn _ fun r ↦ (h r).1, + surjective_of_localized_span _ spn _ fun r ↦ (h r).2⟩ + +lemma exact_of_localized_span + (h : ∀ r : s, Function.Exact (map (.powers r.1) f) (map (.powers r.1) g)) : + Function.Exact f g := + exact_of_isLocalized_span s spn _ (fun _ ↦ mkLinearMap _ _) _ (fun _ ↦ mkLinearMap _ _) + _ (fun _ ↦ mkLinearMap _ _) f g h + +end localized_span diff --git a/Mathlib/RingTheory/LocalProperties/Submodule.lean b/Mathlib/RingTheory/LocalProperties/Submodule.lean index 2ecaa20be887a..90ade41827edf 100644 --- a/Mathlib/RingTheory/LocalProperties/Submodule.lean +++ b/Mathlib/RingTheory/LocalProperties/Submodule.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang, David Swinarski -/ import Mathlib.Algebra.Module.LocalizedModule.Submodule import Mathlib.RingTheory.Localization.AtPrime +import Mathlib.RingTheory.Localization.Away.Basic /-! # Local properties of modules and submodules @@ -80,8 +81,8 @@ theorem Submodule.eq_top_of_localization₀_maximal (N : Submodule R M) 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 + rw [← one_smul R m, ← one_smul R m'] + by_contra 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) @@ -124,3 +125,79 @@ theorem Submodule.eq_top_of_localization_maximal (N : Submodule R M) Submodule.eq_of_localization_maximal Rₚ Mₚ f fun P hP ↦ by simpa using h P end maximal + +section span + +open IsLocalizedModule LocalizedModule Ideal + +variable (s : Set R) (span_eq : Ideal.span s = ⊤) +include span_eq + +variable + (Rₚ : ∀ _ : s, Type*) + [∀ r : s, CommSemiring (Rₚ r)] + [∀ r : s, Algebra R (Rₚ r)] + [∀ r : s, IsLocalization.Away r.1 (Rₚ r)] + (Mₚ : ∀ _ : s, Type*) + [∀ r : s, AddCommMonoid (Mₚ r)] + [∀ r : s, Module R (Mₚ r)] + [∀ r : s, Module (Rₚ r) (Mₚ r)] + [∀ r : s, IsScalarTower R (Rₚ r) (Mₚ r)] + (f : ∀ r : s, M →ₗ[R] Mₚ r) + [∀ r : s, IsLocalizedModule (.powers r.1) (f r)] + +theorem Module.eq_of_isLocalized_span (x y : M) (h : ∀ r : s, f r x = f r y) : x = y := by + suffices Module.eqIdeal R x y = ⊤ by simpa [Module.eqIdeal] using (eq_top_iff_one _).mp this + by_contra ne + have ⟨r, hrs, disj⟩ := exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + have ⟨⟨_, n, rfl⟩, eq⟩ := (IsLocalizedModule.eq_iff_exists (.powers r.1) _).mp (h r) + exact Set.disjoint_left.mp disj eq ⟨n, rfl⟩ + +theorem Module.eq_zero_of_isLocalized_span (x : M) (h : ∀ r : s, f r x = 0) : x = 0 := + eq_of_isLocalized_span s span_eq _ f x 0 <| by simpa only [map_zero] using h + +theorem Submodule.mem_of_isLocalized_span {m : M} {N : Submodule R M} + (h : ∀ r : s, f r m ∈ N.localized₀ (.powers r.1) (f r)) : 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 + by_contra! ne + have ⟨r, hrs, disj⟩ := exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + obtain ⟨a, ha, t, e⟩ := h r + rw [← IsLocalizedModule.mk'_one (.powers r.1), IsLocalizedModule.mk'_eq_mk'_iff] at e + have ⟨u, hu⟩ := e + simp_rw [smul_smul] at hu + exact Set.disjoint_right.mp disj (u * t).2 (by apply hu ▸ smul_mem _ _ ha) + +theorem Submodule.le_of_isLocalized_span {N P : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) ≤ P.localized₀ (.powers r.1) (f r)) : N ≤ P := + fun m hm ↦ mem_of_isLocalized_span s span_eq _ f fun r ↦ h r ⟨m, hm, 1, by simp⟩ + +theorem Submodule.eq_of_isLocalized₀_span {N P : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = P.localized₀ (.powers r.1) (f r)) : N = P := + le_antisymm (le_of_isLocalized_span s span_eq _ _ fun r ↦ (h r).le) + (le_of_isLocalized_span s span_eq _ _ fun r ↦ (h r).ge) + +theorem Submodule.eq_bot_of_isLocalized₀_span {N : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = ⊥) : N = ⊥ := + eq_of_isLocalized₀_span s span_eq Mₚ f fun _ ↦ by simp only [h, Submodule.localized₀_bot] + +theorem Submodule.eq_top_of_isLocalized₀_span {N : Submodule R M} + (h : ∀ r : s, N.localized₀ (.powers r.1) (f r) = ⊤) : N = ⊤ := + eq_of_isLocalized₀_span s span_eq Mₚ f fun _ ↦ by simp only [h, Submodule.localized₀_top] + +theorem Submodule.eq_of_isLocalized'_span {N P : Submodule R M} + (h : ∀ r, N.localized' (Rₚ r) (.powers r.1) (f r) = P.localized' (Rₚ r) (.powers r.1) (f r)) : + N = P := + eq_of_isLocalized₀_span s span_eq _ f fun r ↦ congr(restrictScalars _ $(h r)) + +theorem Submodule.eq_bot_of_isLocalized'_span {N : Submodule R M} + (h : ∀ r : s, N.localized' (Rₚ r) (.powers r.1) (f r) = ⊥) : N = ⊥ := + eq_of_isLocalized'_span s span_eq Rₚ Mₚ f fun _ ↦ by simp only [h, Submodule.localized'_bot] + +theorem Submodule.eq_top_of_isLocalized'_span {N : Submodule R M} + (h : ∀ r : s, N.localized' (Rₚ r) (.powers r.1) (f r) = ⊤) : N = ⊤ := + eq_of_isLocalized'_span s span_eq Rₚ Mₚ f fun _ ↦ by simp only [h, Submodule.localized'_top] + +end span From a3580b7e38bd301d807a41f7be9be153566602bf Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Mon, 25 Nov 2024 11:04:40 +0000 Subject: [PATCH 134/250] chore(Order/ConditionallyCompleteLattice): split off `Defs.lean` from `Basic.lean` (#19277) --- Mathlib.lean | 1 + .../ConditionallyCompleteLattice/Basic.lean | 218 +-------------- .../ConditionallyCompleteLattice/Defs.lean | 251 ++++++++++++++++++ scripts/noshake.json | 2 + 4 files changed, 255 insertions(+), 217 deletions(-) create mode 100644 Mathlib/Order/ConditionallyCompleteLattice/Defs.lean diff --git a/Mathlib.lean b/Mathlib.lean index 2c97f33123aff..34090017c1344 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3918,6 +3918,7 @@ import Mathlib.Order.CompletePartialOrder import Mathlib.Order.CompleteSublattice import Mathlib.Order.Concept import Mathlib.Order.ConditionallyCompleteLattice.Basic +import Mathlib.Order.ConditionallyCompleteLattice.Defs import Mathlib.Order.ConditionallyCompleteLattice.Finset import Mathlib.Order.ConditionallyCompleteLattice.Group import Mathlib.Order.ConditionallyCompleteLattice.Indexed diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean index bf1aa9e9ffaec..437a1e3d52e13 100644 --- a/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean +++ b/Mathlib/Order/ConditionallyCompleteLattice/Basic.lean @@ -3,11 +3,8 @@ Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Sébastien Gouëzel -/ -import Mathlib.Order.Bounds.Basic -import Mathlib.Order.WellFounded -import Mathlib.Data.Set.Image -import Mathlib.Order.Interval.Set.Basic import Mathlib.Data.Set.Lattice +import Mathlib.Order.ConditionallyCompleteLattice.Defs /-! # Theory of conditionally complete lattices @@ -119,50 +116,6 @@ theorem WithBot.coe_sInf' [InfSet α] {s : Set α} (hs : BddBelow s) : end -/-- A conditionally complete lattice is a lattice in which -every nonempty subset which is bounded above has a supremum, and -every nonempty subset which is bounded below has an infimum. -Typical examples are real numbers or natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete lattices, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLattice (α : Type*) extends Lattice α, SupSet α, InfSet α where - /-- `a ≤ sSup s` for all `a ∈ s`. -/ - le_csSup : ∀ s a, BddAbove s → a ∈ s → a ≤ sSup s - /-- `sSup s ≤ a` for all `a ∈ upperBounds s`. -/ - csSup_le : ∀ s a, Set.Nonempty s → a ∈ upperBounds s → sSup s ≤ a - /-- `sInf s ≤ a` for all `a ∈ s`. -/ - csInf_le : ∀ s a, BddBelow s → a ∈ s → sInf s ≤ a - /-- `a ≤ sInf s` for all `a ∈ lowerBounds s`. -/ - le_csInf : ∀ s a, Set.Nonempty s → a ∈ lowerBounds s → a ≤ sInf s - --- Porting note: mathlib3 used `renaming` -/-- A conditionally complete linear order is a linear order in which -every nonempty subset which is bounded above has a supremum, and -every nonempty subset which is bounded below has an infimum. -Typical examples are real numbers or natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLinearOrder (α : Type*) extends ConditionallyCompleteLattice α where - /-- A `ConditionallyCompleteLinearOrder` is total. -/ - le_total (a b : α) : a ≤ b ∨ b ≤ a - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableLE : DecidableRel (· ≤ · : α → α → Prop) - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE - /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ - decidableLT : DecidableRel (· < · : α → α → Prop) := - @decidableLTOfDecidableLE _ _ decidableLE - /-- If a set is not bounded above, its supremum is by convention `sSup ∅`. -/ - csSup_of_not_bddAbove : ∀ s, ¬BddAbove s → sSup s = sSup (∅ : Set α) - /-- If a set is not bounded below, its infimum is by convention `sInf ∅`. -/ - csInf_of_not_bddBelow : ∀ s, ¬BddBelow s → sInf s = sInf (∅ : Set α) - instance ConditionallyCompleteLinearOrder.toLinearOrder [ConditionallyCompleteLinearOrder α] : LinearOrder α := { ‹ConditionallyCompleteLinearOrder α› with @@ -179,19 +132,6 @@ instance ConditionallyCompleteLinearOrder.toLinearOrder [ConditionallyCompleteLi · simp [h₁] · simp [show ¬(a ≤ b) from fun h => hab (le_antisymm h h₂), h₂] } -/-- A conditionally complete linear order with `Bot` is a linear order with least element, in which -every nonempty subset which is bounded above has a supremum, and every nonempty subset (necessarily -bounded below) has an infimum. A typical example is the natural numbers. - -To differentiate the statements from the corresponding statements in (unconditional) -complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should -hold in both worlds, sometimes with additional assumptions of nonemptiness or -boundedness. -/ -class ConditionallyCompleteLinearOrderBot (α : Type*) extends ConditionallyCompleteLinearOrder α, - OrderBot α where - /-- The supremum of the empty set is special-cased to `⊥` -/ - csSup_empty : sSup ∅ = ⊥ - -- see Note [lower instance priority] attribute [instance 100] ConditionallyCompleteLinearOrderBot.toOrderBot @@ -214,36 +154,6 @@ instance (priority := 100) CompleteLinearOrder.toConditionallyCompleteLinearOrde csSup_of_not_bddAbove := fun s H ↦ (H (OrderTop.bddAbove s)).elim csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } -open scoped Classical in -/-- A well founded linear order is conditionally complete, with a bottom element. -/ -noncomputable abbrev WellFoundedLT.conditionallyCompleteLinearOrderBot (α : Type*) - [i₁ : LinearOrder α] [i₂ : OrderBot α] [h : WellFoundedLT α] : - ConditionallyCompleteLinearOrderBot α := - { i₁, i₂, LinearOrder.toLattice with - sInf := fun s => if hs : s.Nonempty then h.wf.min s hs else ⊥ - csInf_le := fun s a _ has => by - have s_ne : s.Nonempty := ⟨a, has⟩ - simpa [s_ne] using not_lt.1 (h.wf.not_lt_min s s_ne has) - le_csInf := fun s a hs has => by - simp only [hs, dif_pos] - exact has (h.wf.min_mem s hs) - sSup := fun s => if hs : (upperBounds s).Nonempty then h.wf.min _ hs else ⊥ - le_csSup := fun s a hs has => by - have h's : (upperBounds s).Nonempty := hs - simp only [h's, dif_pos] - exact h.wf.min_mem _ h's has - csSup_le := fun s a _ has => by - have h's : (upperBounds s).Nonempty := ⟨a, has⟩ - simp only [h's, dif_pos] - simpa using h.wf.not_lt_min _ h's has - csSup_empty := by simpa using eq_bot_iff.2 (not_lt.1 <| h.wf.not_lt_min _ _ <| mem_univ ⊥) - csSup_of_not_bddAbove := by - intro s H - have B : ¬((upperBounds s).Nonempty) := H - simp only [B, dite_false, upperBounds_empty, univ_nonempty, dite_true] - exact le_antisymm bot_le (WellFounded.min_le _ (mem_univ _)) - csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } - namespace OrderDual instance instConditionallyCompleteLattice (α : Type*) [ConditionallyCompleteLattice α] : @@ -261,132 +171,6 @@ instance (α : Type*) [ConditionallyCompleteLinearOrder α] : ConditionallyCompl end OrderDual -/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `sup` function -that returns the least upper bound of a nonempty set which is bounded above. Usually this -constructor provides poor definitional equalities. If other fields are known explicitly, they -should be provided; for example, if `inf` is known explicitly, construct the -`ConditionallyCompleteLattice` instance as -``` -instance : ConditionallyCompleteLattice my_T := - { inf := better_inf, - le_inf := ..., - inf_le_right := ..., - inf_le_left := ... - -- don't care to fix sup, sInf - ..conditionallyCompleteLatticeOfsSup my_T _ } -``` --/ -def conditionallyCompleteLatticeOfsSup (α : Type*) [H1 : PartialOrder α] [H2 : SupSet α] - (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) - (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) - (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : - ConditionallyCompleteLattice α := - { H1, H2 with - sup := fun a b => sSup {a, b} - le_sup_left := fun a b => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) - le_sup_right := fun a b => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 - (mem_insert_of_mem _ (mem_singleton _)) - sup_le := fun a b _ hac hbc => - (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).2 - (forall_insert_of_forall (forall_eq.mpr hbc) hac) - inf := fun a b => sSup (lowerBounds {a, b}) - inf_le_left := fun a b => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - (bddBelow_pair a b)).2 - fun _ hc => hc <| mem_insert _ _ - inf_le_right := fun a b => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - (bddBelow_pair a b)).2 - fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) - le_inf := fun c a b hca hcb => - (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) - ⟨c, forall_insert_of_forall (forall_eq.mpr hcb) hca⟩).1 - (forall_insert_of_forall (forall_eq.mpr hcb) hca) - sInf := fun s => sSup (lowerBounds s) - csSup_le := fun s a hs ha => (isLUB_sSup s ⟨a, ha⟩ hs).2 ha - le_csSup := fun s a hs ha => (isLUB_sSup s hs ⟨a, ha⟩).1 ha - csInf_le := fun s a hs ha => - (isLUB_sSup (lowerBounds s) (Nonempty.bddAbove_lowerBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha - le_csInf := fun s a hs ha => - (isLUB_sSup (lowerBounds s) hs.bddAbove_lowerBounds ⟨a, ha⟩).1 ha } - -/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `inf` function -that returns the greatest lower bound of a nonempty set which is bounded below. Usually this -constructor provides poor definitional equalities. If other fields are known explicitly, they -should be provided; for example, if `inf` is known explicitly, construct the -`ConditionallyCompleteLattice` instance as -``` -instance : ConditionallyCompleteLattice my_T := - { inf := better_inf, - le_inf := ..., - inf_le_right := ..., - inf_le_left := ... - -- don't care to fix sup, sSup - ..conditionallyCompleteLatticeOfsInf my_T _ } -``` --/ -def conditionallyCompleteLatticeOfsInf (α : Type*) [H1 : PartialOrder α] [H2 : InfSet α] - (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) - (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) - (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : - ConditionallyCompleteLattice α := - { H1, H2 with - inf := fun a b => sInf {a, b} - inf_le_left := fun a b => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) - inf_le_right := fun a b => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 - (mem_insert_of_mem _ (mem_singleton _)) - le_inf := fun _ a b hca hcb => - (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).2 - (forall_insert_of_forall (forall_eq.mpr hcb) hca) - sup := fun a b => sInf (upperBounds {a, b}) - le_sup_left := fun a b => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - (bddAbove_pair a b)).2 - fun _ hc => hc <| mem_insert _ _ - le_sup_right := fun a b => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - (bddAbove_pair a b)).2 - fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) - sup_le := fun a b c hac hbc => - (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) - ⟨c, forall_insert_of_forall (forall_eq.mpr hbc) hac⟩).1 - (forall_insert_of_forall (forall_eq.mpr hbc) hac) - sSup := fun s => sInf (upperBounds s) - le_csInf := fun s a hs ha => (isGLB_sInf s ⟨a, ha⟩ hs).2 ha - csInf_le := fun s a hs ha => (isGLB_sInf s hs ⟨a, ha⟩).1 ha - le_csSup := fun s a hs ha => - (isGLB_sInf (upperBounds s) (Nonempty.bddBelow_upperBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha - csSup_le := fun s a hs ha => - (isGLB_sInf (upperBounds s) hs.bddBelow_upperBounds ⟨a, ha⟩).1 ha } - -/-- A version of `conditionallyCompleteLatticeOfsSup` when we already know that `α` is a lattice. - -This should only be used when it is both hard and unnecessary to provide `inf` explicitly. -/ -def conditionallyCompleteLatticeOfLatticeOfsSup (α : Type*) [H1 : Lattice α] [SupSet α] - (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : - ConditionallyCompleteLattice α := - { H1, - conditionallyCompleteLatticeOfsSup α - (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) - (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) - isLUB_sSup with } - -/-- A version of `conditionallyCompleteLatticeOfsInf` when we already know that `α` is a lattice. - -This should only be used when it is both hard and unnecessary to provide `sup` explicitly. -/ -def conditionallyCompleteLatticeOfLatticeOfsInf (α : Type*) [H1 : Lattice α] [InfSet α] - (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : - ConditionallyCompleteLattice α := - { H1, - conditionallyCompleteLatticeOfsInf α - (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) - (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) - isGLB_sInf with } - section ConditionallyCompleteLattice variable [ConditionallyCompleteLattice α] {s t : Set α} {a b : α} diff --git a/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean b/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean new file mode 100644 index 0000000000000..4a499bbdfdc7a --- /dev/null +++ b/Mathlib/Order/ConditionallyCompleteLattice/Defs.lean @@ -0,0 +1,251 @@ +/- +Copyright (c) 2018 Sébastien Gouëzel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Sébastien Gouëzel +-/ +import Mathlib.Order.Bounds.Basic +import Mathlib.Order.SetNotation +import Mathlib.Order.WellFounded + +/-! +# Definitions of conditionally complete lattices + +A conditionally complete lattice is a lattice in which every non-empty bounded subset `s` +has a least upper bound and a greatest lower bound, denoted below by `sSup s` and `sInf s`. +Typical examples are `ℝ`, `ℕ`, and `ℤ` with their usual orders. + +The theory is very comparable to the theory of complete lattices, except that suitable +boundedness and nonemptiness assumptions have to be added to most statements. +We express these using the `BddAbove` and `BddBelow` predicates, which we use to prove +most useful properties of `sSup` and `sInf` in conditionally complete lattices. + +To differentiate the statements between complete lattices and conditionally complete +lattices, we prefix `sInf` and `sSup` in the statements by `c`, giving `csInf` and `csSup`. +For instance, `sInf_le` is a statement in complete lattices ensuring `sInf s ≤ x`, +while `csInf_le` is the same statement in conditionally complete lattices +with an additional assumption that `s` is bounded below. +-/ + +open Set + +variable {α β γ : Type*} {ι : Sort*} + +/-- A conditionally complete lattice is a lattice in which +every nonempty subset which is bounded above has a supremum, and +every nonempty subset which is bounded below has an infimum. +Typical examples are real numbers or natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete lattices, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLattice (α : Type*) extends Lattice α, SupSet α, InfSet α where + /-- `a ≤ sSup s` for all `a ∈ s`. -/ + le_csSup : ∀ s a, BddAbove s → a ∈ s → a ≤ sSup s + /-- `sSup s ≤ a` for all `a ∈ upperBounds s`. -/ + csSup_le : ∀ s a, Set.Nonempty s → a ∈ upperBounds s → sSup s ≤ a + /-- `sInf s ≤ a` for all `a ∈ s`. -/ + csInf_le : ∀ s a, BddBelow s → a ∈ s → sInf s ≤ a + /-- `a ≤ sInf s` for all `a ∈ lowerBounds s`. -/ + le_csInf : ∀ s a, Set.Nonempty s → a ∈ lowerBounds s → a ≤ sInf s + +-- Porting note: mathlib3 used `renaming` +/-- A conditionally complete linear order is a linear order in which +every nonempty subset which is bounded above has a supremum, and +every nonempty subset which is bounded below has an infimum. +Typical examples are real numbers or natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLinearOrder (α : Type*) extends ConditionallyCompleteLattice α where + /-- A `ConditionallyCompleteLinearOrder` is total. -/ + le_total (a b : α) : a ≤ b ∨ b ≤ a + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableLE : DecidableRel (· ≤ · : α → α → Prop) + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE + /-- In a `ConditionallyCompleteLinearOrder`, we assume the order relations are all decidable. -/ + decidableLT : DecidableRel (· < · : α → α → Prop) := + @decidableLTOfDecidableLE _ _ decidableLE + /-- If a set is not bounded above, its supremum is by convention `sSup ∅`. -/ + csSup_of_not_bddAbove : ∀ s, ¬BddAbove s → sSup s = sSup (∅ : Set α) + /-- If a set is not bounded below, its infimum is by convention `sInf ∅`. -/ + csInf_of_not_bddBelow : ∀ s, ¬BddBelow s → sInf s = sInf (∅ : Set α) + +/-- A conditionally complete linear order with `Bot` is a linear order with least element, in which +every nonempty subset which is bounded above has a supremum, and every nonempty subset (necessarily +bounded below) has an infimum. A typical example is the natural numbers. + +To differentiate the statements from the corresponding statements in (unconditional) +complete linear orders, we prefix `sInf` and `sSup` by a `c` everywhere. The same statements should +hold in both worlds, sometimes with additional assumptions of nonemptiness or +boundedness. -/ +class ConditionallyCompleteLinearOrderBot (α : Type*) extends ConditionallyCompleteLinearOrder α, + OrderBot α where + /-- The supremum of the empty set is special-cased to `⊥` -/ + csSup_empty : sSup ∅ = ⊥ + +-- see Note [lower instance priority] +attribute [instance 100] ConditionallyCompleteLinearOrderBot.toOrderBot + +open scoped Classical in +/-- A well founded linear order is conditionally complete, with a bottom element. -/ +noncomputable abbrev WellFoundedLT.conditionallyCompleteLinearOrderBot (α : Type*) + [i₁ : LinearOrder α] [i₂ : OrderBot α] [h : WellFoundedLT α] : + ConditionallyCompleteLinearOrderBot α := + { i₁, i₂, LinearOrder.toLattice with + sInf := fun s => if hs : s.Nonempty then h.wf.min s hs else ⊥ + csInf_le := fun s a _ has => by + have s_ne : s.Nonempty := ⟨a, has⟩ + simpa [s_ne] using not_lt.1 (h.wf.not_lt_min s s_ne has) + le_csInf := fun s a hs has => by + simp only [hs, dif_pos] + exact has (h.wf.min_mem s hs) + sSup := fun s => if hs : (upperBounds s).Nonempty then h.wf.min _ hs else ⊥ + le_csSup := fun s a hs has => by + have h's : (upperBounds s).Nonempty := hs + simp only [h's, dif_pos] + exact h.wf.min_mem _ h's has + csSup_le := fun s a _ has => by + have h's : (upperBounds s).Nonempty := ⟨a, has⟩ + simp only [h's, dif_pos] + simpa using h.wf.not_lt_min _ h's has + csSup_empty := by simpa using eq_bot_iff.2 (not_lt.1 <| h.wf.not_lt_min _ _ <| mem_univ ⊥) + csSup_of_not_bddAbove := by + intro s H + have B : ¬((upperBounds s).Nonempty) := H + simp only [B, dite_false, upperBounds_empty, univ_nonempty, dite_true] + exact le_antisymm bot_le (WellFounded.min_le _ (mem_univ _)) + csInf_of_not_bddBelow := fun s H ↦ (H (OrderBot.bddBelow s)).elim } + +namespace OrderDual + +end OrderDual + +/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `sup` function +that returns the least upper bound of a nonempty set which is bounded above. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`ConditionallyCompleteLattice` instance as +``` +instance : ConditionallyCompleteLattice my_T := + { inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, sInf + ..conditionallyCompleteLatticeOfsSup my_T _ } +``` +-/ +def conditionallyCompleteLatticeOfsSup (α : Type*) [H1 : PartialOrder α] [H2 : SupSet α] + (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) + (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) + (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : + ConditionallyCompleteLattice α := + { H1, H2 with + sup := fun a b => sSup {a, b} + le_sup_left := fun a b => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) + le_sup_right := fun a b => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).1 + (mem_insert_of_mem _ (mem_singleton _)) + sup_le := fun a b _ hac hbc => + (isLUB_sSup {a, b} (bddAbove_pair a b) (insert_nonempty _ _)).2 + (forall_insert_of_forall (forall_eq.mpr hbc) hac) + inf := fun a b => sSup (lowerBounds {a, b}) + inf_le_left := fun a b => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + (bddBelow_pair a b)).2 + fun _ hc => hc <| mem_insert _ _ + inf_le_right := fun a b => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + (bddBelow_pair a b)).2 + fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) + le_inf := fun c a b hca hcb => + (isLUB_sSup (lowerBounds {a, b}) (Nonempty.bddAbove_lowerBounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hcb) hca⟩).1 + (forall_insert_of_forall (forall_eq.mpr hcb) hca) + sInf := fun s => sSup (lowerBounds s) + csSup_le := fun s a hs ha => (isLUB_sSup s ⟨a, ha⟩ hs).2 ha + le_csSup := fun s a hs ha => (isLUB_sSup s hs ⟨a, ha⟩).1 ha + csInf_le := fun s a hs ha => + (isLUB_sSup (lowerBounds s) (Nonempty.bddAbove_lowerBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha + le_csInf := fun s a hs ha => + (isLUB_sSup (lowerBounds s) hs.bddAbove_lowerBounds ⟨a, ha⟩).1 ha } + +/-- Create a `ConditionallyCompleteLattice` from a `PartialOrder` and `inf` function +that returns the greatest lower bound of a nonempty set which is bounded below. Usually this +constructor provides poor definitional equalities. If other fields are known explicitly, they +should be provided; for example, if `inf` is known explicitly, construct the +`ConditionallyCompleteLattice` instance as +``` +instance : ConditionallyCompleteLattice my_T := + { inf := better_inf, + le_inf := ..., + inf_le_right := ..., + inf_le_left := ... + -- don't care to fix sup, sSup + ..conditionallyCompleteLatticeOfsInf my_T _ } +``` +-/ +def conditionallyCompleteLatticeOfsInf (α : Type*) [H1 : PartialOrder α] [H2 : InfSet α] + (bddAbove_pair : ∀ a b : α, BddAbove ({a, b} : Set α)) + (bddBelow_pair : ∀ a b : α, BddBelow ({a, b} : Set α)) + (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : + ConditionallyCompleteLattice α := + { H1, H2 with + inf := fun a b => sInf {a, b} + inf_le_left := fun a b => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 (mem_insert _ _) + inf_le_right := fun a b => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).1 + (mem_insert_of_mem _ (mem_singleton _)) + le_inf := fun _ a b hca hcb => + (isGLB_sInf {a, b} (bddBelow_pair a b) (insert_nonempty _ _)).2 + (forall_insert_of_forall (forall_eq.mpr hcb) hca) + sup := fun a b => sInf (upperBounds {a, b}) + le_sup_left := fun a b => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + (bddAbove_pair a b)).2 + fun _ hc => hc <| mem_insert _ _ + le_sup_right := fun a b => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + (bddAbove_pair a b)).2 + fun _ hc => hc <| mem_insert_of_mem _ (mem_singleton _) + sup_le := fun a b c hac hbc => + (isGLB_sInf (upperBounds {a, b}) (Nonempty.bddBelow_upperBounds ⟨a, mem_insert _ _⟩) + ⟨c, forall_insert_of_forall (forall_eq.mpr hbc) hac⟩).1 + (forall_insert_of_forall (forall_eq.mpr hbc) hac) + sSup := fun s => sInf (upperBounds s) + le_csInf := fun s a hs ha => (isGLB_sInf s ⟨a, ha⟩ hs).2 ha + csInf_le := fun s a hs ha => (isGLB_sInf s hs ⟨a, ha⟩).1 ha + le_csSup := fun s a hs ha => + (isGLB_sInf (upperBounds s) (Nonempty.bddBelow_upperBounds ⟨a, ha⟩) hs).2 fun _ hb => hb ha + csSup_le := fun s a hs ha => + (isGLB_sInf (upperBounds s) hs.bddBelow_upperBounds ⟨a, ha⟩).1 ha } + +/-- A version of `conditionallyCompleteLatticeOfsSup` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `inf` explicitly. -/ +def conditionallyCompleteLatticeOfLatticeOfsSup (α : Type*) [H1 : Lattice α] [SupSet α] + (isLUB_sSup : ∀ s : Set α, BddAbove s → s.Nonempty → IsLUB s (sSup s)) : + ConditionallyCompleteLattice α := + { H1, + conditionallyCompleteLatticeOfsSup α + (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + isLUB_sSup with } + +/-- A version of `conditionallyCompleteLatticeOfsInf` when we already know that `α` is a lattice. + +This should only be used when it is both hard and unnecessary to provide `sup` explicitly. -/ +def conditionallyCompleteLatticeOfLatticeOfsInf (α : Type*) [H1 : Lattice α] [InfSet α] + (isGLB_sInf : ∀ s : Set α, BddBelow s → s.Nonempty → IsGLB s (sInf s)) : + ConditionallyCompleteLattice α := + { H1, + conditionallyCompleteLatticeOfsInf α + (fun a b => ⟨a ⊔ b, forall_insert_of_forall (forall_eq.mpr le_sup_right) le_sup_left⟩) + (fun a b => ⟨a ⊓ b, forall_insert_of_forall (forall_eq.mpr inf_le_right) inf_le_left⟩) + isGLB_sInf with } diff --git a/scripts/noshake.json b/scripts/noshake.json index 86cfab061a740..4a06c0c7680d4 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -385,6 +385,8 @@ ["Mathlib.Control.Basic", "Mathlib.Data.Option.Basic"], "Mathlib.Data.LazyList.Basic": ["Mathlib.Lean.Thunk"], "Mathlib.Data.Int.Defs": ["Batteries.Data.Int.Order"], + "Mathlib.Data.Int.ConditionallyCompleteOrder": + ["Mathlib.Order.ConditionallyCompleteLattice.Basic"], "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"], From c4f307606793b52847c977499c58d745bc2186d3 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 11:04:41 +0000 Subject: [PATCH 135/250] chore: remove unused variables (#19413) #17715 --- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 2 +- Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean | 2 +- Mathlib/Algebra/Homology/Additive.lean | 4 ++-- Mathlib/Algebra/Quaternion.lean | 4 ++-- Mathlib/Algebra/TrivSqZeroExt.lean | 5 ++--- Mathlib/AlgebraicGeometry/Morphisms/Basic.lean | 5 ++--- Mathlib/AlgebraicGeometry/Morphisms/Etale.lean | 2 +- Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean | 2 +- Mathlib/AlgebraicGeometry/SpreadingOut.lean | 3 +-- Mathlib/Analysis/Convex/Gauge.lean | 2 +- Mathlib/Analysis/Convex/Normed.lean | 4 ++-- Mathlib/Analysis/InnerProductSpace/Basic.lean | 10 +++------- Mathlib/Analysis/Normed/Field/Basic.lean | 4 ++-- Mathlib/Analysis/Normed/Group/Basic.lean | 8 ++++---- Mathlib/Analysis/Normed/Group/Bounded.lean | 3 +-- Mathlib/Analysis/Normed/Group/Uniform.lean | 8 ++++---- Mathlib/Analysis/Normed/Module/FiniteDimension.lean | 7 ++----- Mathlib/CategoryTheory/Adjunction/Evaluation.lean | 1 - Mathlib/CategoryTheory/GradedObject.lean | 2 +- Mathlib/CategoryTheory/GuitartExact/Basic.lean | 1 - .../CategoryTheory/Sites/NonabelianCohomology/H1.lean | 2 +- Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean | 3 +-- .../Combinatorics/SimpleGraph/Triangle/Counting.lean | 4 ++-- .../Combinatorics/SimpleGraph/Triangle/Tripartite.lean | 2 +- Mathlib/Data/Real/Sqrt.lean | 2 -- Mathlib/MeasureTheory/Constructions/Pi.lean | 2 +- Mathlib/MeasureTheory/Group/Action.lean | 4 ++-- Mathlib/MeasureTheory/Group/LIntegral.lean | 2 +- Mathlib/MeasureTheory/Integral/Marginal.lean | 2 +- Mathlib/MeasureTheory/MeasurableSpace/Basic.lean | 4 ++-- Mathlib/MeasureTheory/MeasurableSpace/Prod.lean | 2 +- Mathlib/MeasureTheory/Measure/Regular.lean | 4 ++-- .../GradedAlgebra/HomogeneousLocalization.lean | 3 +-- Mathlib/RingTheory/Localization/AtPrime.lean | 2 +- Mathlib/RingTheory/Localization/Integral.lean | 2 +- .../Localization/LocalizationLocalization.lean | 5 ++--- Mathlib/RingTheory/TensorProduct/Free.lean | 2 +- Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean | 2 -- Mathlib/Testing/Plausible/Functions.lean | 4 ++-- Mathlib/Topology/Algebra/InfiniteSum/Basic.lean | 6 +++--- .../Topology/Algebra/InfiniteSum/Constructions.lean | 2 +- Mathlib/Topology/Algebra/InfiniteSum/Defs.lean | 2 +- Mathlib/Topology/Algebra/InfiniteSum/Group.lean | 4 ++-- Mathlib/Topology/Algebra/InfiniteSum/Ring.lean | 9 ++++----- Mathlib/Topology/ContinuousMap/Ordered.lean | 3 +-- Mathlib/Topology/Instances/ENNReal.lean | 2 +- Mathlib/Topology/MetricSpace/Dilation.lean | 6 +++--- Mathlib/Topology/MetricSpace/Isometry.lean | 2 +- Mathlib/Topology/MetricSpace/Pseudo/Real.lean | 4 ++-- Mathlib/Topology/Order/ExtrClosure.lean | 2 +- Mathlib/Topology/Order/MonotoneConvergence.lean | 8 ++++---- Mathlib/Topology/Order/T5.lean | 3 +-- Mathlib/Topology/ShrinkingLemma.lean | 2 +- Mathlib/Topology/UniformSpace/OfCompactT2.lean | 2 +- 54 files changed, 83 insertions(+), 106 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 9472500e7fcc9..404fdd46e25ad 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -38,7 +38,7 @@ lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by end Mul section InvolutiveInv -variable [InvolutiveInv G] {s t : Set G} +variable [InvolutiveInv G] @[to_additive (attr := simp)] lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean index 8466c8876f1c1..4e144b113a816 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Set/Card.lean @@ -13,7 +13,7 @@ import Mathlib.SetTheory.Cardinal.Finite open scoped Cardinal Pointwise -variable {G G₀ M M₀ : Type*} +variable {G₀ M₀ : Type*} namespace Set variable [GroupWithZero G₀] [Zero M₀] [MulActionWithZero G₀ M₀] {a : G₀} diff --git a/Mathlib/Algebra/Homology/Additive.lean b/Mathlib/Algebra/Homology/Additive.lean index 5fbebb2c8beb9..1b14c45bf5b00 100644 --- a/Mathlib/Algebra/Homology/Additive.lean +++ b/Mathlib/Algebra/Homology/Additive.lean @@ -23,8 +23,8 @@ variable {ι : Type*} variable {V : Type u} [Category.{v} V] [Preadditive V] variable {W : Type*} [Category W] [Preadditive W] variable {W₁ W₂ : Type*} [Category W₁] [Category W₂] [HasZeroMorphisms W₁] [HasZeroMorphisms W₂] -variable {c : ComplexShape ι} {C D E : HomologicalComplex V c} -variable (f g : C ⟶ D) (h k : D ⟶ E) (i : ι) +variable {c : ComplexShape ι} {C D : HomologicalComplex V c} +variable (f : C ⟶ D) (i : ι) namespace HomologicalComplex diff --git a/Mathlib/Algebra/Quaternion.lean b/Mathlib/Algebra/Quaternion.lean index 69ce497aff36a..1ee91c3592cc2 100644 --- a/Mathlib/Algebra/Quaternion.lean +++ b/Mathlib/Algebra/Quaternion.lean @@ -93,7 +93,7 @@ theorem equivTuple_apply {R : Type*} (c₁ c₂ : R) (x : ℍ[R,c₁,c₂]) : @[simp] theorem mk.eta {R : Type*} {c₁ c₂} (a : ℍ[R,c₁,c₂]) : mk a.1 a.2 a.3 a.4 = a := rfl -variable {S T R : Type*} {c₁ c₂ : R} (r x y z : R) (a b c : ℍ[R,c₁,c₂]) +variable {S T R : Type*} {c₁ c₂ : R} (r x y : R) (a b : ℍ[R,c₁,c₂]) instance [Subsingleton R] : Subsingleton ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).subsingleton instance [Nontrivial R] : Nontrivial ℍ[R, c₁, c₂] := (equivTuple c₁ c₂).surjective.nontrivial @@ -742,7 +742,7 @@ instance {R : Type*} [One R] [Neg R] [Nontrivial R] : Nontrivial ℍ[R] := namespace Quaternion -variable {S T R : Type*} [CommRing R] (r x y z : R) (a b c : ℍ[R]) +variable {S T R : Type*} [CommRing R] (r x y : R) (a b : ℍ[R]) export QuaternionAlgebra (re imI imJ imK) diff --git a/Mathlib/Algebra/TrivSqZeroExt.lean b/Mathlib/Algebra/TrivSqZeroExt.lean index c40ab975c18d7..cc4239ab85e25 100644 --- a/Mathlib/Algebra/TrivSqZeroExt.lean +++ b/Mathlib/Algebra/TrivSqZeroExt.lean @@ -870,10 +870,9 @@ section Algebra variable (S : Type*) (R R' : Type u) (M : Type v) variable [CommSemiring S] [Semiring R] [CommSemiring R'] [AddCommMonoid M] -variable [Algebra S R] [Algebra S R'] [Module S M] -variable [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] +variable [Algebra S R] [Module S M] [Module R M] [Module Rᵐᵒᵖ M] [SMulCommClass R Rᵐᵒᵖ M] variable [IsScalarTower S R M] [IsScalarTower S Rᵐᵒᵖ M] -variable [Module R' M] [Module R'ᵐᵒᵖ M] [IsCentralScalar R' M] [IsScalarTower S R' M] +variable [Module R' M] [Module R'ᵐᵒᵖ M] [IsCentralScalar R' M] instance algebra' : Algebra S (tsze R M) := { (TrivSqZeroExt.inlHom R M).comp (algebraMap S R) with diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean index 51f9c3ccf41d9..070c60d31cc77 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Basic.lean @@ -142,8 +142,7 @@ instance inf (P Q : MorphismProperty Scheme) [IsLocalAtTarget P] [IsLocalAtTarge fun h ↦ ⟨(iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).left), (iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).right)⟩⟩ -variable {P} [hP : IsLocalAtTarget P] -variable {X Y U V : Scheme.{u}} {f : X ⟶ Y} {g : U ⟶ Y} [IsOpenImmersion g] (𝒰 : Y.OpenCover) +variable {P} [hP : IsLocalAtTarget P] {X Y : Scheme.{u}} {f : X ⟶ Y} (𝒰 : Y.OpenCover) lemma of_isPullback {UX UY : Scheme.{u}} {iY : UY ⟶ Y} [IsOpenImmersion iY] {iX : UX ⟶ X} {f' : UX ⟶ UY} (h : IsPullback iX f' f iY) (H : P f) : P f' := by @@ -228,7 +227,7 @@ instance inf (P Q : MorphismProperty Scheme) [IsLocalAtSource P] [IsLocalAtSourc (iff_of_openCover' f 𝒰).mpr (fun i ↦ (h i).right)⟩⟩ variable {P} [IsLocalAtSource P] -variable {X Y U V : Scheme.{u}} {f : X ⟶ Y} {g : U ⟶ Y} [IsOpenImmersion g] (𝒰 : X.OpenCover) +variable {X Y : Scheme.{u}} {f : X ⟶ Y} (𝒰 : X.OpenCover) lemma comp {UX : Scheme.{u}} (H : P f) (i : UX ⟶ X) [IsOpenImmersion i] : P (i ≫ f) := diff --git a/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean b/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean index 665513b4b09cb..35ee3280f7055 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/Etale.lean @@ -28,7 +28,7 @@ abbrev IsEtale {X Y : Scheme.{u}} (f : X ⟶ Y) := IsSmoothOfRelativeDimension 0 namespace IsEtale -variable {X Y : Scheme.{u}} (f : X ⟶ Y) +variable {X : Scheme.{u}} instance : IsStableUnderBaseChange @IsEtale := isSmoothOfRelativeDimension_isStableUnderBaseChange 0 diff --git a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean index 880a1a62832f1..6cd62e8973f1b 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/OpenImmersion.lean @@ -26,7 +26,7 @@ universe u namespace AlgebraicGeometry -variable {X Y Z : Scheme.{u}} (f : X ⟶ Y) (g : Y ⟶ Z) +variable {X Y : Scheme.{u}} theorem isOpenImmersion_iff_stalk {f : X ⟶ Y} : IsOpenImmersion f ↔ IsOpenEmbedding f.base ∧ ∀ x, IsIso (f.stalkMap x) := by diff --git a/Mathlib/AlgebraicGeometry/SpreadingOut.lean b/Mathlib/AlgebraicGeometry/SpreadingOut.lean index b5d7ef83b97b8..a327e1ae7ee02 100644 --- a/Mathlib/AlgebraicGeometry/SpreadingOut.lean +++ b/Mathlib/AlgebraicGeometry/SpreadingOut.lean @@ -49,8 +49,7 @@ open CategoryTheory namespace AlgebraicGeometry -variable {X Y S : Scheme.{u}} (f : X ⟶ Y) (sX : X ⟶ S) (sY : Y ⟶ S) (e : f ≫ sY = sX) -variable {R A : CommRingCat.{u}} +variable {X Y S : Scheme.{u}} (f : X ⟶ Y) (sX : X ⟶ S) (sY : Y ⟶ S) {R A : CommRingCat.{u}} /-- The germ map at `x` is injective if there exists some affine `U ∋ x` such that the map `Γ(X, U) ⟶ X_x` is injective -/ diff --git a/Mathlib/Analysis/Convex/Gauge.lean b/Mathlib/Analysis/Convex/Gauge.lean index 5ecb5e8b2c4cd..f9d037d23877a 100644 --- a/Mathlib/Analysis/Convex/Gauge.lean +++ b/Mathlib/Analysis/Convex/Gauge.lean @@ -41,7 +41,7 @@ open scoped Pointwise Topology NNReal noncomputable section -variable {𝕜 E F : Type*} +variable {𝕜 E : Type*} section AddCommGroup diff --git a/Mathlib/Analysis/Convex/Normed.lean b/Mathlib/Analysis/Convex/Normed.lean index fae3e58b55f18..9a305c1457de0 100644 --- a/Mathlib/Analysis/Convex/Normed.lean +++ b/Mathlib/Analysis/Convex/Normed.lean @@ -25,14 +25,14 @@ We prove the following facts: is bounded. -/ -variable {ι : Type*} {E P : Type*} +variable {E P : Type*} open AffineBasis Module Metric Set open scoped Convex Pointwise Topology section SeminormedAddCommGroup variable [SeminormedAddCommGroup E] [NormedSpace ℝ E] [PseudoMetricSpace P] [NormedAddTorsor E P] -variable {s t : Set E} +variable {s : Set E} /-- The norm on a real normed space is convex on any convex set. See also `Seminorm.convexOn` and `convexOn_univ_norm`. -/ diff --git a/Mathlib/Analysis/InnerProductSpace/Basic.lean b/Mathlib/Analysis/InnerProductSpace/Basic.lean index 2f715eb17ed19..6f06b27cb078c 100644 --- a/Mathlib/Analysis/InnerProductSpace/Basic.lean +++ b/Mathlib/Analysis/InnerProductSpace/Basic.lean @@ -1600,9 +1600,7 @@ end ContinuousLinearMap section -variable {ι : Type*} {ι' : Type*} {ι'' : Type*} -variable {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] -variable {E'' : Type*} [SeminormedAddCommGroup E''] [InnerProductSpace 𝕜 E''] +variable {ι : Type*} {ι' : Type*} {E' : Type*} [SeminormedAddCommGroup E'] [InnerProductSpace 𝕜 E'] @[simp] theorem Orthonormal.equiv_refl {v : Basis ι 𝕜 E} (hv : Orthonormal 𝕜 v) : @@ -1683,7 +1681,7 @@ open scoped InnerProductSpace variable [NormedAddCommGroup E] [InnerProductSpace 𝕜 E] variable [NormedAddCommGroup F] [InnerProductSpace ℝ F] -variable {ι : Type*} {ι' : Type*} {ι'' : Type*} +variable {ι : Type*} local notation "⟪" x ", " y "⟫" => @inner 𝕜 _ _ x y @@ -2188,9 +2186,7 @@ local notation "IK" => @RCLike.I 𝕜 _ local postfix:90 "†" => starRingEnd _ -variable {ι : Type*} -variable {G : ι → Type*} [∀ i, NormedAddCommGroup (G i)] [∀ i, InnerProductSpace 𝕜 (G i)] - {V : ∀ i, G i →ₗᵢ[𝕜] E} (hV : OrthogonalFamily 𝕜 G V) [dec_V : ∀ (i) (x : G i), Decidable (x ≠ 0)] +variable {ι : Type*} {G : ι → Type*} /-- An orthogonal family forms an independent family of subspaces; that is, any collection of elements each from a different subspace in the family is linearly independent. In particular, the diff --git a/Mathlib/Analysis/Normed/Field/Basic.lean b/Mathlib/Analysis/Normed/Field/Basic.lean index eb3c22b61051f..00424d49f8973 100644 --- a/Mathlib/Analysis/Normed/Field/Basic.lean +++ b/Mathlib/Analysis/Normed/Field/Basic.lean @@ -955,7 +955,7 @@ theorem NormedAddCommGroup.tendsto_atTop' [Nonempty α] [Preorder α] [IsDirecte section RingHomIsometric -variable {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} +variable {R₁ R₂ : Type*} /-- This class states that a ring homomorphism is isometric. This is a sufficient assumption for a continuous semilinear map to be bounded and this is the main use for this typeclass. -/ @@ -965,7 +965,7 @@ class RingHomIsometric [Semiring R₁] [Semiring R₂] [Norm R₁] [Norm R₂] ( attribute [simp] RingHomIsometric.is_iso -variable [SeminormedRing R₁] [SeminormedRing R₂] [SeminormedRing R₃] +variable [SeminormedRing R₁] instance RingHomIsometric.ids : RingHomIsometric (RingHom.id R₁) := ⟨rfl⟩ diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index 91cdbe19d79a5..dd3bf2a5b2797 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -44,7 +44,7 @@ normed group -/ -variable {𝓕 𝕜 α ι κ E F G : Type*} +variable {𝓕 α ι κ E F G : Type*} open Filter Function Metric Bornology @@ -337,7 +337,7 @@ abbrev GroupNorm.toNormedCommGroup [CommGroup E] (f : GroupNorm E) : NormedCommG section SeminormedGroup variable [SeminormedGroup E] [SeminormedGroup F] [SeminormedGroup G] {s : Set E} - {a a₁ a₂ b b₁ b₂ c : E} {r r₁ r₂ : ℝ} + {a a₁ a₂ b c : E} {r r₁ r₂ : ℝ} @[to_additive] theorem dist_eq_norm_div (a b : E) : dist a b = ‖a / b‖ := @@ -1024,7 +1024,7 @@ end Induced section SeminormedCommGroup -variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a b : E} {r : ℝ} @[to_additive] theorem dist_inv (x y : E) : dist x⁻¹ y = dist x y⁻¹ := by @@ -1290,7 +1290,7 @@ end SeminormedCommGroup section NormedGroup -variable [NormedGroup E] [NormedGroup F] {a b : E} +variable [NormedGroup E] {a b : E} @[to_additive (attr := simp) norm_eq_zero] theorem norm_eq_zero'' : ‖a‖ = 0 ↔ a = 1 := diff --git a/Mathlib/Analysis/Normed/Group/Bounded.lean b/Mathlib/Analysis/Normed/Group/Bounded.lean index bd28a0b52f340..dcebdaec8bfb2 100644 --- a/Mathlib/Analysis/Normed/Group/Bounded.lean +++ b/Mathlib/Analysis/Normed/Group/Bounded.lean @@ -21,11 +21,10 @@ normed group open Filter Metric Bornology open scoped Pointwise Topology -variable {α ι E F G : Type*} +variable {α E F G : Type*} section SeminormedGroup variable [SeminormedGroup E] [SeminormedGroup F] [SeminormedGroup G] {s : Set E} - {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} @[to_additive (attr := simp) comap_norm_atTop] lemma comap_norm_atTop' : comap norm atTop = cobounded E := by diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 8f4a7f9f06b4e..73b33458e9210 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -15,13 +15,13 @@ This file proves lipschitzness of normed group operations and shows that normed groups. -/ -variable {𝓕 α E F : Type*} +variable {𝓕 E F : Type*} open Filter Function Metric Bornology open scoped ENNReal NNReal Uniformity Pointwise Topology section SeminormedGroup -variable [SeminormedGroup E] [SeminormedGroup F] {s : Set E} {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedGroup E] [SeminormedGroup F] {s : Set E} {a b : E} {r : ℝ} @[to_additive] instance NormedGroup.to_isometricSMul_right : IsometricSMul Eᵐᵒᵖ E := @@ -177,7 +177,7 @@ end SeminormedGroup section SeminormedCommGroup -variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a a₁ a₂ b b₁ b₂ : E} {r r₁ r₂ : ℝ} +variable [SeminormedCommGroup E] [SeminormedCommGroup F] {a₁ a₂ b₁ b₂ : E} {r₁ r₂ : ℝ} @[to_additive] instance NormedGroup.to_isometricSMul_left : IsometricSMul E E := @@ -239,7 +239,7 @@ theorem edist_mul_mul_le (a₁ a₂ b₁ b₂ : E) : section PseudoEMetricSpace variable {α E : Type*} [SeminormedCommGroup E] [PseudoEMetricSpace α] {K Kf Kg : ℝ≥0} - {f g : α → E} {s : Set α} {x : α} + {f g : α → E} {s : Set α} @[to_additive (attr := simp)] lemma lipschitzWith_inv_iff : LipschitzWith K f⁻¹ ↔ LipschitzWith K f := by simp [LipschitzWith] diff --git a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean index 303247548cea2..789c690f2042f 100644 --- a/Mathlib/Analysis/Normed/Module/FiniteDimension.lean +++ b/Mathlib/Analysis/Normed/Module/FiniteDimension.lean @@ -54,8 +54,7 @@ namespace LinearIsometry open LinearMap -variable {R : Type*} [Semiring R] -variable {F E₁ : Type*} [SeminormedAddCommGroup F] [NormedAddCommGroup E₁] [Module R E₁] +variable {F E₁ : Type*} [SeminormedAddCommGroup F] [NormedAddCommGroup E₁] variable {R₁ : Type*} [Field R₁] [Module R₁ E₁] [Module R₁ F] [FiniteDimensional R₁ E₁] [FiniteDimensional R₁ F] @@ -110,9 +109,7 @@ end AffineIsometry section CompleteField variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type v} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {F : Type w} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {F' : Type x} - [AddCommGroup F'] [Module 𝕜 F'] [TopologicalSpace F'] [TopologicalAddGroup F'] - [ContinuousSMul 𝕜 F'] [CompleteSpace 𝕜] + [NormedSpace 𝕜 E] {F : Type w} [NormedAddCommGroup F] [NormedSpace 𝕜 F] [CompleteSpace 𝕜] section Affine diff --git a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean index d11a3d4ac62c1..ffb59cad14a86 100644 --- a/Mathlib/CategoryTheory/Adjunction/Evaluation.lean +++ b/Mathlib/CategoryTheory/Adjunction/Evaluation.lean @@ -22,7 +22,6 @@ open CategoryTheory.Limits universe v₁ v₂ v₃ u₁ u₂ u₃ variable {C : Type u₁} [Category.{v₁} C] (D : Type u₂) [Category.{v₂} D] - (E : Type u₃) [Category.{v₃} E] noncomputable section diff --git a/Mathlib/CategoryTheory/GradedObject.lean b/Mathlib/CategoryTheory/GradedObject.lean index b6e2df434c2b9..2bc160aeba95d 100644 --- a/Mathlib/CategoryTheory/GradedObject.lean +++ b/Mathlib/CategoryTheory/GradedObject.lean @@ -316,7 +316,7 @@ lemma hasMap_of_iso (e : X ≅ Y) (p: I → J) [HasMap X p] : HasMap Y p := fun exact hasColimitOfIso α.symm section -variable [X.HasMap p] [Y.HasMap p] [Z.HasMap p] +variable [X.HasMap p] [Y.HasMap p] /-- Given `X : GradedObject I C` and `p : I → J`, `X.mapObj p` is the graded object by `J` which in degree `j` consists of the coproduct of the `X i` such that `p i = j`. -/ diff --git a/Mathlib/CategoryTheory/GuitartExact/Basic.lean b/Mathlib/CategoryTheory/GuitartExact/Basic.lean index cb64500850d99..b697c1ee0f1a1 100644 --- a/Mathlib/CategoryTheory/GuitartExact/Basic.lean +++ b/Mathlib/CategoryTheory/GuitartExact/Basic.lean @@ -128,7 +128,6 @@ abbrev CostructuredArrowDownwards.mk (comm : R.map a ≫ w.app X₁ ≫ B.map b CostructuredArrow.mk (Y := StructuredArrow.mk a) (StructuredArrow.homMk b (by simpa using comm)) -variable (comm : R.map a ≫ w.app X₁ ≫ B.map b = g) variable {w g} lemma StructuredArrowRightwards.mk_surjective diff --git a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean index 44a948a9222b1..b15c8b4f9117f 100644 --- a/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean +++ b/Mathlib/CategoryTheory/Sites/NonabelianCohomology/H1.lean @@ -48,7 +48,7 @@ variable {C : Type u} [Category.{v} C] namespace PresheafOfGroups -variable (G : Cᵒᵖ ⥤ Grp.{w}) {X : C} {I : Type w'} (U : I → C) +variable (G : Cᵒᵖ ⥤ Grp.{w}) {I : Type w'} (U : I → C) /-- A zero cochain consists of a family of sections. -/ def ZeroCochain := ∀ (i : I), G.obj (Opposite.op (U i)) diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean index 6d675a7865092..2924a54b4d530 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Basic.lean @@ -34,8 +34,7 @@ open Fintype (card) namespace SimpleGraph -variable {α β 𝕜 : Type*} [LinearOrderedField 𝕜] {G H : SimpleGraph α} {ε δ : 𝕜} {n : ℕ} - {s : Finset α} +variable {α β 𝕜 : Type*} [LinearOrderedField 𝕜] {G H : SimpleGraph α} {ε δ : 𝕜} section LocallyLinear diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean index bff996c3fa38c..3de98cd5194b3 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Counting.lean @@ -23,7 +23,7 @@ attribute [-instance] decidableEq_of_subsingleton open Finset Fintype -variable {α : Type*} (G G' : SimpleGraph α) [DecidableRel G.Adj] {ε : ℝ} {s t u : Finset α} +variable {α : Type*} (G : SimpleGraph α) [DecidableRel G.Adj] {ε : ℝ} {s t u : Finset α} namespace SimpleGraph @@ -146,7 +146,7 @@ private lemma triple_eq_triple_of_mem (hst : Disjoint s t) (hsu : Disjoint s u) · rintro rfl solve_by_elim -variable [Fintype α] {P : Finpartition (univ : Finset α)} +variable [Fintype α] /-- The **Triangle Counting Lemma**. If `G` is a graph and `s`, `t`, `u` are disjoint sets of vertices such that each pair is `ε`-uniform and `2 * ε`-dense, then `G` contains at least diff --git a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean index 3135218de5e3d..c911166bc53ac 100644 --- a/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean +++ b/Mathlib/Combinatorics/SimpleGraph/Triangle/Tripartite.lean @@ -38,7 +38,7 @@ This construction shows up unrelatedly twice in the theory of Roth numbers: open Finset Function Sum3 variable {α β γ 𝕜 : Type*} [LinearOrderedField 𝕜] {t : Finset (α × β × γ)} {a a' : α} {b b' : β} - {c c' : γ} {x : α × β × γ} {ε : 𝕜} + {c c' : γ} {x : α × β × γ} namespace SimpleGraph namespace TripartiteFromTriangles diff --git a/Mathlib/Data/Real/Sqrt.lean b/Mathlib/Data/Real/Sqrt.lean index b43bd46488f69..5b1193b70b404 100644 --- a/Mathlib/Data/Real/Sqrt.lean +++ b/Mathlib/Data/Real/Sqrt.lean @@ -311,8 +311,6 @@ end Mathlib.Meta.Positivity namespace Real -variable {x y : ℝ} - @[simp] theorem sqrt_mul {x : ℝ} (hx : 0 ≤ x) (y : ℝ) : √(x * y) = √x * √y := by simp_rw [Real.sqrt, ← NNReal.coe_mul, NNReal.coe_inj, Real.toNNReal_mul hx, NNReal.sqrt_mul] diff --git a/Mathlib/MeasureTheory/Constructions/Pi.lean b/Mathlib/MeasureTheory/Constructions/Pi.lean index aafd0f1b3c5ef..670e4e5bf0347 100644 --- a/Mathlib/MeasureTheory/Constructions/Pi.lean +++ b/Mathlib/MeasureTheory/Constructions/Pi.lean @@ -80,7 +80,7 @@ theorem isPiSystem_pi [∀ i, MeasurableSpace (α i)] : section Finite -variable [Finite ι] [Finite ι'] +variable [Finite ι] /-- Boxes of countably spanning sets are countably spanning. -/ theorem IsCountablySpanning.pi {C : ∀ i, Set (Set (α i))} (hC : ∀ i, IsCountablySpanning (C i)) : diff --git a/Mathlib/MeasureTheory/Group/Action.lean b/Mathlib/MeasureTheory/Group/Action.lean index 0cb212b5273b1..cad802e003d1f 100644 --- a/Mathlib/MeasureTheory/Group/Action.lean +++ b/Mathlib/MeasureTheory/Group/Action.lean @@ -28,7 +28,7 @@ namespace MeasureTheory universe u v w -variable {G : Type u} {M : Type v} {α : Type w} {s : Set α} +variable {G : Type u} {M : Type v} {α : Type w} namespace SMulInvariantMeasure @@ -198,7 +198,7 @@ instance smulInvariantMeasure_map_smul [SMul M α] [SMul N α] [SMulCommClass N end SMulHomClass -variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] (c : G) (μ : Measure α) +variable (G) {m : MeasurableSpace α} [Group G] [MulAction G α] (μ : Measure α) variable [MeasurableSpace G] [MeasurableSMul G α] in /-- Equivalent definitions of a measure invariant under a multiplicative action of a group. diff --git a/Mathlib/MeasureTheory/Group/LIntegral.lean b/Mathlib/MeasureTheory/Group/LIntegral.lean index e0344acdeed93..74c8453b516af 100644 --- a/Mathlib/MeasureTheory/Group/LIntegral.lean +++ b/Mathlib/MeasureTheory/Group/LIntegral.lean @@ -20,7 +20,7 @@ open Measure TopologicalSpace open scoped ENNReal -variable {G : Type*} [MeasurableSpace G] {μ : Measure G} {g : G} +variable {G : Type*} [MeasurableSpace G] {μ : Measure G} section MeasurableMul diff --git a/Mathlib/MeasureTheory/Integral/Marginal.lean b/Mathlib/MeasureTheory/Integral/Marginal.lean index 76fcaa682df92..4e19390aa37fb 100644 --- a/Mathlib/MeasureTheory/Integral/Marginal.lean +++ b/Mathlib/MeasureTheory/Integral/Marginal.lean @@ -66,7 +66,7 @@ section LMarginal variable {δ δ' : Type*} {π : δ → Type*} [∀ x, MeasurableSpace (π x)] variable {μ : ∀ i, Measure (π i)} [DecidableEq δ] -variable {s t : Finset δ} {f g : (∀ i, π i) → ℝ≥0∞} {x y : ∀ i, π i} {i : δ} +variable {s t : Finset δ} {f : (∀ i, π i) → ℝ≥0∞} {x : ∀ i, π i} /-- Integrate `f(x₁,…,xₙ)` over all variables `xᵢ` where `i ∈ s`. Return a function in the remaining variables (it will be constant in the `xᵢ` for `i ∈ s`). diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean index 4d2eb7fd31ab1..678a6983571d0 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Basic.lean @@ -55,7 +55,7 @@ open Set Encodable Function Equiv Filter MeasureTheory universe uι -variable {α β γ δ δ' : Type*} {ι : Sort uι} {s t u : Set α} +variable {α β γ δ δ' : Type*} {ι : Sort uι} {s : Set α} namespace MeasurableSpace @@ -225,7 +225,7 @@ variable {f g : α → β} section TypeclassMeasurableSpace -variable [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] +variable [MeasurableSpace α] [MeasurableSpace β] @[nontriviality, measurability] theorem Subsingleton.measurable [Subsingleton α] : Measurable f := fun _ _ => diff --git a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean index e9f2db14bfc25..9dec6c3e39e9b 100644 --- a/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean +++ b/Mathlib/MeasureTheory/MeasurableSpace/Prod.lean @@ -18,7 +18,7 @@ noncomputable section open Set MeasurableSpace -variable {α β γ : Type*} [MeasurableSpace α] [MeasurableSpace β] [MeasurableSpace γ] +variable {α β : Type*} [MeasurableSpace α] [MeasurableSpace β] /-- The product of generated σ-algebras is the one generated by rectangles, if both generating sets are countably spanning. -/ diff --git a/Mathlib/MeasureTheory/Measure/Regular.lean b/Mathlib/MeasureTheory/Measure/Regular.lean index ed27aa6936522..91b924e0a4f30 100644 --- a/Mathlib/MeasureTheory/Measure/Regular.lean +++ b/Mathlib/MeasureTheory/Measure/Regular.lean @@ -446,7 +446,7 @@ protected theorem FiniteSpanningSetsIn.outerRegular namespace InnerRegularWRT -variable {p q : Set α → Prop} {U s : Set α} {ε r : ℝ≥0∞} +variable {p : Set α → Prop} /-- If the restrictions of a measure to a monotone sequence of sets covering the space are inner regular for some property `p` and all measurable sets, then the measure itself is @@ -637,7 +637,7 @@ end InnerRegularWRT namespace InnerRegular -variable {U : Set α} {ε : ℝ≥0∞} [TopologicalSpace α] +variable [TopologicalSpace α] /-- The measure of a measurable set is the supremum of the measures of compact sets it contains. -/ theorem _root_.MeasurableSet.measure_eq_iSup_isCompact ⦃U : Set α⦄ (hU : MeasurableSet U) diff --git a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean index af6dc235b4928..82a5c54c0a72c 100644 --- a/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean +++ b/Mathlib/RingTheory/GradedAlgebra/HomogeneousLocalization.lean @@ -636,8 +636,7 @@ 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 {e : ι} {f : A} {g : A} (hg : g ∈ 𝒜 e) {x : A} (hx : x = f * g) variable (𝒜) diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 65cccdc214ca9..76d9a9d2cb2e8 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -31,7 +31,7 @@ commutative ring, field of fractions -/ -variable {R : Type*} [CommSemiring R] (M : Submonoid R) (S : Type*) [CommSemiring S] +variable {R : Type*} [CommSemiring R] (S : Type*) [CommSemiring S] variable [Algebra R S] {P : Type*} [CommSemiring P] section AtPrime diff --git a/Mathlib/RingTheory/Localization/Integral.lean b/Mathlib/RingTheory/Localization/Integral.lean index 07c55c2708e4b..ad327225f2383 100644 --- a/Mathlib/RingTheory/Localization/Integral.lean +++ b/Mathlib/RingTheory/Localization/Integral.lean @@ -25,7 +25,7 @@ commutative ring, field of fractions variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] -variable [Algebra R S] {P : Type*} [CommRing P] +variable [Algebra R S] open Polynomial diff --git a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean index 62db6f2a12761..0637b7cd2c470 100644 --- a/Mathlib/RingTheory/Localization/LocalizationLocalization.lean +++ b/Mathlib/RingTheory/Localization/LocalizationLocalization.lean @@ -27,8 +27,7 @@ namespace IsLocalization section LocalizationLocalization -variable {R : Type*} [CommSemiring R] (M : Submonoid R) {S : Type*} [CommSemiring S] -variable [Algebra R S] {P : Type*} [CommSemiring P] +variable {R : Type*} [CommSemiring R] (M : Submonoid R) {S : Type*} [CommSemiring S] [Algebra R S] variable (N : Submonoid S) (T : Type*) [CommSemiring T] [Algebra R T] @@ -251,7 +250,7 @@ end IsLocalization namespace IsFractionRing -variable {R : Type*} [CommRing R] (M : Submonoid R) {S : Type*} [CommRing S] +variable {R : Type*} [CommRing R] (M : Submonoid R) open IsLocalization diff --git a/Mathlib/RingTheory/TensorProduct/Free.lean b/Mathlib/RingTheory/TensorProduct/Free.lean index 1f644e47cf568..a857cfdf874c6 100644 --- a/Mathlib/RingTheory/TensorProduct/Free.lean +++ b/Mathlib/RingTheory/TensorProduct/Free.lean @@ -30,7 +30,7 @@ namespace Algebra namespace TensorProduct -variable {R S A : Type*} +variable {R A : Type*} section Basis diff --git a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean index 680aa8f4249fe..4c69bfb655b75 100644 --- a/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean +++ b/Mathlib/RingTheory/UniqueFactorizationDomain/Nat.lean @@ -15,8 +15,6 @@ import Mathlib.RingTheory.UniqueFactorizationDomain.NormalizedFactors * `Nat.instUniqueFactorizationMonoid`: the natural numbers have unique factorization -/ -variable {α : Type*} - namespace Nat instance instWfDvdMonoid : WfDvdMonoid ℕ where diff --git a/Mathlib/Testing/Plausible/Functions.lean b/Mathlib/Testing/Plausible/Functions.lean index 9a83323198f47..a9bdf24d8ccb4 100644 --- a/Mathlib/Testing/Plausible/Functions.lean +++ b/Mathlib/Testing/Plausible/Functions.lean @@ -40,9 +40,9 @@ their defining property is invariant through shrinking. Injective functions are an example of how complicated it can get. -/ -universe u v w +universe u v -variable {α : Type u} {β : Type v} {γ : Sort w} +variable {α : Type u} {β : Type v} namespace Plausible diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean index 74529c1335d72..ba49c204bab72 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Basic.lean @@ -21,12 +21,12 @@ noncomputable section open Filter Finset Function Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} section HasProd variable [CommMonoid α] [TopologicalSpace α] -variable {f g : β → α} {a b : α} {s : Finset β} +variable {f g : β → α} {a b : α} /-- Constant one function has product `1` -/ @[to_additive "Constant zero function has sum `0`"] @@ -355,7 +355,7 @@ end HasProd section tprod -variable [CommMonoid α] [TopologicalSpace α] {f g : β → α} {a a₁ a₂ : α} +variable [CommMonoid α] [TopologicalSpace α] {f g : β → α} @[to_additive] theorem tprod_congr_set_coe (f : β → α) {s t : Set β} (h : s = t) : diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean index 8b1f1a092440e..4386fe90f6bd1 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Constructions.lean @@ -20,7 +20,7 @@ open Filter Finset Function open scoped Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} /-! ## Product, Sigma and Pi types -/ diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index 2bf685cd201e1..cc83507487809 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -111,7 +111,7 @@ notation3 "∏' "(...)", "r:67:(scoped f => tprod f) => r @[inherit_doc tsum] notation3 "∑' "(...)", "r:67:(scoped f => tsum f) => r -variable {f g : β → α} {a b : α} {s : Finset β} +variable {f : β → α} {a : α} {s : Finset β} @[to_additive] theorem HasProd.multipliable (h : HasProd f a) : Multipliable f := diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean index 4c5353916aae3..937a415544928 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Group.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Group.lean @@ -20,7 +20,7 @@ open Filter Finset Function open scoped Topology -variable {α β γ δ : Type*} +variable {α β γ : Type*} section TopologicalGroup @@ -197,7 +197,7 @@ theorem multipliable_iff_cauchySeq_finset [CompleteSpace α] {f : β → α} : Multipliable f ↔ CauchySeq fun s : Finset β ↦ ∏ b ∈ s, f b := by classical exact cauchy_map_iff_exists_tendsto.symm -variable [UniformGroup α] {f g : β → α} {a a₁ a₂ : α} +variable [UniformGroup α] {f g : β → α} @[to_additive] theorem cauchySeq_finset_iff_prod_vanishing : diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean index ea55550f54743..aa0cdf8f9a57d 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Ring.lean @@ -19,12 +19,12 @@ This file provides lemmas about the interaction between infinite sums and multip open Filter Finset Function -variable {ι κ R α : Type*} +variable {ι κ α : Type*} section NonUnitalNonAssocSemiring -variable [NonUnitalNonAssocSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α} - {a a₁ a₂ : α} +variable [NonUnitalNonAssocSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f : ι → α} + {a₁ : α} theorem HasSum.mul_left (a₂) (h : HasSum f a₁) : HasSum (fun i ↦ a₂ * f i) (a₂ * a₁) := by simpa only using h.map (AddMonoidHom.mulLeft a₂) (continuous_const.mul continuous_id) @@ -63,8 +63,7 @@ end NonUnitalNonAssocSemiring section DivisionSemiring -variable [DivisionSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f g : ι → α} - {a a₁ a₂ : α} +variable [DivisionSemiring α] [TopologicalSpace α] [TopologicalSemiring α] {f : ι → α} {a a₁ a₂ : α} theorem HasSum.div_const (h : HasSum f a) (b : α) : HasSum (fun i ↦ f i / b) (a / b) := by simp only [div_eq_mul_inv, h.mul_right b⁻¹] diff --git a/Mathlib/Topology/ContinuousMap/Ordered.lean b/Mathlib/Topology/ContinuousMap/Ordered.lean index 5ef9876bb5f8c..7677b1c20ea56 100644 --- a/Mathlib/Topology/ContinuousMap/Ordered.lean +++ b/Mathlib/Topology/ContinuousMap/Ordered.lean @@ -13,8 +13,7 @@ import Mathlib.Topology.ContinuousMap.Defs -/ -variable {α : Type*} {β : Type*} {γ : Type*} -variable [TopologicalSpace α] [TopologicalSpace β] [TopologicalSpace γ] +variable {α β : Type*} [TopologicalSpace α] [TopologicalSpace β] namespace ContinuousMap diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 94b73607b7f71..029ada7f6a78a 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -25,7 +25,7 @@ variable {α : Type*} {β : Type*} {γ : Type*} namespace ENNReal -variable {a b c d : ℝ≥0∞} {r p q : ℝ≥0} {x y z : ℝ≥0∞} {ε ε₁ ε₂ : ℝ≥0∞} {s : Set ℝ≥0∞} +variable {a b : ℝ≥0∞} {r : ℝ≥0} {x : ℝ≥0∞} {ε : ℝ≥0∞} section TopologicalSpace diff --git a/Mathlib/Topology/MetricSpace/Dilation.lean b/Mathlib/Topology/MetricSpace/Dilation.lean index 201c0e06799e3..6ab93147d696b 100644 --- a/Mathlib/Topology/MetricSpace/Dilation.lean +++ b/Mathlib/Topology/MetricSpace/Dilation.lean @@ -74,7 +74,7 @@ end Defs namespace Dilation -variable {α : Type*} {β : Type*} {γ : Type*} {F : Type*} {G : Type*} +variable {α : Type*} {β : Type*} {γ : Type*} {F : Type*} section Setup @@ -229,8 +229,8 @@ end Setup section PseudoEmetricDilation variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] [PseudoEMetricSpace γ] -variable [FunLike F α β] [DilationClass F α β] [FunLike G β γ] [DilationClass G β γ] -variable (f : F) (g : G) {x y z : α} {s : Set α} +variable [FunLike F α β] [DilationClass F α β] +variable (f : F) /-- Every isometry is a dilation of ratio `1`. -/ @[simps] diff --git a/Mathlib/Topology/MetricSpace/Isometry.lean b/Mathlib/Topology/MetricSpace/Isometry.lean index 5e434b37434df..a987f5c09f284 100644 --- a/Mathlib/Topology/MetricSpace/Isometry.lean +++ b/Mathlib/Topology/MetricSpace/Isometry.lean @@ -62,7 +62,7 @@ namespace Isometry section PseudoEmetricIsometry variable [PseudoEMetricSpace α] [PseudoEMetricSpace β] [PseudoEMetricSpace γ] -variable {f : α → β} {x y z : α} {s : Set α} +variable {f : α → β} {x : α} /-- An isometry preserves edistances. -/ theorem edist_eq (hf : Isometry f) (x y : α) : edist (f x) (f y) = edist x y := diff --git a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean index baa4f534f75cc..6183ec868a947 100644 --- a/Mathlib/Topology/MetricSpace/Pseudo/Real.lean +++ b/Mathlib/Topology/MetricSpace/Pseudo/Real.lean @@ -15,7 +15,7 @@ open scoped NNReal Topology namespace Real -variable {ι α : Type*} [PseudoMetricSpace α] +variable {ι : Type*} lemma dist_left_le_of_mem_uIcc {x y z : ℝ} (h : y ∈ uIcc x z) : dist x y ≤ dist x z := by simpa only [dist_comm x] using abs_sub_left_of_mem_uIcc h @@ -35,7 +35,7 @@ lemma dist_le_of_mem_Icc {x y x' y' : ℝ} (hx : x ∈ Icc x' y') (hy : y ∈ Ic lemma dist_le_of_mem_Icc_01 {x y : ℝ} (hx : x ∈ Icc (0 : ℝ) 1) (hy : y ∈ Icc (0 : ℝ) 1) : dist x y ≤ 1 := by simpa only [sub_zero] using Real.dist_le_of_mem_Icc hx hy -variable {π : ι → Type*} [Fintype ι] [∀ i, PseudoMetricSpace (π i)] {x y x' y' : ι → ℝ} +variable [Fintype ι] {x y x' y' : ι → ℝ} lemma dist_le_of_mem_pi_Icc (hx : x ∈ Icc x' y') (hy : y ∈ Icc x' y') : dist x y ≤ dist x' y' := by refine (dist_pi_le_iff dist_nonneg).2 fun b => diff --git a/Mathlib/Topology/Order/ExtrClosure.lean b/Mathlib/Topology/Order/ExtrClosure.lean index d1b2049789057..31b06abf781bd 100644 --- a/Mathlib/Topology/Order/ExtrClosure.lean +++ b/Mathlib/Topology/Order/ExtrClosure.lean @@ -20,7 +20,7 @@ open Filter Set open Topology variable {X Y : Type*} [TopologicalSpace X] [TopologicalSpace Y] [Preorder Y] - [OrderClosedTopology Y] {f g : X → Y} {s : Set X} {a : X} + [OrderClosedTopology Y] {f : X → Y} {s : Set X} {a : X} protected theorem IsMaxOn.closure (h : IsMaxOn f s a) (hc : ContinuousOn f (closure s)) : IsMaxOn f (closure s) a := fun x hx => diff --git a/Mathlib/Topology/Order/MonotoneConvergence.lean b/Mathlib/Topology/Order/MonotoneConvergence.lean index bad7ee58b291e..7936b81bae9f3 100644 --- a/Mathlib/Topology/Order/MonotoneConvergence.lean +++ b/Mathlib/Topology/Order/MonotoneConvergence.lean @@ -107,7 +107,7 @@ end IsGLB section CiSup -variable [ConditionallyCompleteLattice α] [SupConvergenceClass α] {f : ι → α} {a : α} +variable [ConditionallyCompleteLattice α] [SupConvergenceClass α] {f : ι → α} theorem tendsto_atTop_ciSup (h_mono : Monotone f) (hbdd : BddAbove <| range f) : Tendsto f atTop (𝓝 (⨆ i, f i)) := by @@ -121,7 +121,7 @@ end CiSup section CiInf -variable [ConditionallyCompleteLattice α] [InfConvergenceClass α] {f : ι → α} {a : α} +variable [ConditionallyCompleteLattice α] [InfConvergenceClass α] {f : ι → α} theorem tendsto_atBot_ciInf (h_mono : Monotone f) (hbdd : BddBelow <| range f) : Tendsto f atBot (𝓝 (⨅ i, f i)) := by convert tendsto_atTop_ciSup h_mono.dual hbdd.dual using 1 @@ -133,7 +133,7 @@ end CiInf section iSup -variable [CompleteLattice α] [SupConvergenceClass α] {f : ι → α} {a : α} +variable [CompleteLattice α] [SupConvergenceClass α] {f : ι → α} theorem tendsto_atTop_iSup (h_mono : Monotone f) : Tendsto f atTop (𝓝 (⨆ i, f i)) := tendsto_atTop_ciSup h_mono (OrderTop.bddAbove _) @@ -145,7 +145,7 @@ end iSup section iInf -variable [CompleteLattice α] [InfConvergenceClass α] {f : ι → α} {a : α} +variable [CompleteLattice α] [InfConvergenceClass α] {f : ι → α} theorem tendsto_atBot_iInf (h_mono : Monotone f) : Tendsto f atBot (𝓝 (⨅ i, f i)) := tendsto_atBot_ciInf h_mono (OrderBot.bddBelow _) diff --git a/Mathlib/Topology/Order/T5.lean b/Mathlib/Topology/Order/T5.lean index 39e0d08dc9b0a..597085db3f109 100644 --- a/Mathlib/Topology/Order/T5.lean +++ b/Mathlib/Topology/Order/T5.lean @@ -16,8 +16,7 @@ topological space. open Filter Set Function OrderDual Topology Interval -variable {X : Type*} [LinearOrder X] [TopologicalSpace X] [OrderTopology X] {a b c : X} - {s t : Set X} +variable {X : Type*} [LinearOrder X] [TopologicalSpace X] [OrderTopology X] {a : X} {s t : Set X} namespace Set diff --git a/Mathlib/Topology/ShrinkingLemma.lean b/Mathlib/Topology/ShrinkingLemma.lean index 8f03eeab6a068..65f6b72889730 100644 --- a/Mathlib/Topology/ShrinkingLemma.lean +++ b/Mathlib/Topology/ShrinkingLemma.lean @@ -206,7 +206,7 @@ section NormalSpace open ShrinkingLemma -variable {u : ι → Set X} {s : Set X} {p : Set X → Prop} [NormalSpace X] +variable {u : ι → Set X} {s : Set X} [NormalSpace X] /-- **Shrinking lemma**. A point-finite open cover of a closed subset of a normal space can be "shrunk" to a new open cover so that the closure of each new open set is contained in the diff --git a/Mathlib/Topology/UniformSpace/OfCompactT2.lean b/Mathlib/Topology/UniformSpace/OfCompactT2.lean index 5eaf71a852f4f..c9d49e9474570 100644 --- a/Mathlib/Topology/UniformSpace/OfCompactT2.lean +++ b/Mathlib/Topology/UniformSpace/OfCompactT2.lean @@ -26,7 +26,7 @@ uniform space, uniform continuity, compact space open Uniformity Topology Filter UniformSpace Set -variable {α β γ : Type*} [UniformSpace α] [UniformSpace β] +variable {γ : Type*} /-! ### Uniformity on compact spaces From 2fa86db9f80c9a25fe911ff8ba831094a2659078 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Mon, 25 Nov 2024 11:42:16 +0000 Subject: [PATCH 136/250] feat(AlgebraicTopology): StrictSegal simplicial sets are quasicategories (#19270) We prove that any simplicial set satisfying the `StrictSegal` condition introduced in #18499 is a quasicategory. This implies that the nerve of a category is a quasicategory. We construct a solution to the lifting problem against `hornInclusion n i` as a `spineToSimplex` of the path in the image of the spine of the standard simplex. To prove that this extension restricts to the appropriate map on the horn, we apply `spineInjective`, reducing the question of equality between two `n`-simplices in `X` to that of equality between the paths along their spines. The remainder of the proof is a straightforward application of basic properties relating spines of simplices to their faces. Co-Authored-By: [Johan Commelin](https://github.com/jcommelin) and [Emily Riehl](https://github.com/emilyriehl) --- .../AlgebraicTopology/SimplexCategory.lean | 44 ++++++ .../SimplicialSet/Basic.lean | 7 + .../AlgebraicTopology/SimplicialSet/Path.lean | 50 ++++++ .../SimplicialSet/StrictSegal.lean | 148 +++++++++++++++++- 4 files changed, 248 insertions(+), 1 deletion(-) diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index a2e26b4118880..7c9ce0278a0c3 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -564,6 +564,50 @@ lemma δ_one_mkOfSucc {n : ℕ} (i : Fin n) : fin_cases x aesop +/-- If `i + 1 < j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i` and `i + 1`, respectively. -/ +lemma mkOfSucc_δ_lt {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : i.succ.castSucc < j) : + mkOfSucc i ≫ δ j = mkOfSucc i.castSucc := by + ext x + fin_cases x + · simp [δ, Fin.succAbove_of_castSucc_lt _ _ (Nat.lt_trans _ h)] + · simp [δ, Fin.succAbove_of_castSucc_lt _ _ h] + +/-- If `i + 1 > j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i + 1` and `i + 2`, respectively. -/ +lemma mkOfSucc_δ_gt {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : j < i.succ.castSucc) : + mkOfSucc i ≫ δ j = mkOfSucc i.succ := by + ext x + simp only [δ, len_mk, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, OrderHom.comp_coe, + OrderEmbedding.toOrderHom_coe, Function.comp_apply, Fin.succAboveOrderEmb_apply] + fin_cases x <;> rw [Fin.succAbove_of_le_castSucc] + · rfl + · exact Nat.le_of_lt_succ h + · rfl + · exact Nat.le_of_lt h + +/-- If `i + 1 = j`, `mkOfSucc i ≫ δ j` is the morphism `[1] ⟶ [n]` that +sends `0` and `1` to `i` and `i + 2`, respectively. -/ +lemma mkOfSucc_δ_eq {n : ℕ} {i : Fin n} {j : Fin (n + 2)} + (h : j = i.succ.castSucc) : + mkOfSucc i ≫ δ j = intervalEdge i 2 (by omega) := by + ext x + fin_cases x + · subst h + simp only [δ, len_mk, Nat.reduceAdd, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, + Fin.zero_eta, OrderHom.comp_coe, OrderEmbedding.toOrderHom_coe, Function.comp_apply, + mkOfSucc_homToOrderHom_zero, Fin.succAboveOrderEmb_apply, + Fin.castSucc_succAbove_castSucc, Fin.succAbove_succ_self] + rfl + · simp only [δ, len_mk, Nat.reduceAdd, mkHom, comp_toOrderHom, Hom.toOrderHom_mk, Fin.mk_one, + OrderHom.comp_coe, OrderEmbedding.toOrderHom_coe, Function.comp_apply, + mkOfSucc_homToOrderHom_one, Fin.succAboveOrderEmb_apply] + subst h + rw [Fin.succAbove_castSucc_self] + rfl + theorem eq_of_one_to_two (f : ([1] : SimplexCategory) ⟶ [2]) : f = (δ (n := 1) 0) ∨ f = (δ (n := 1) 1) ∨ f = (δ (n := 1) 2) ∨ ∃ a, f = SimplexCategory.const _ _ a := by diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 309fc78b1916a..8c329680dec88 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -106,6 +106,13 @@ def _root_.SSet.yonedaEquiv (X : SSet.{u}) (n : SimplexCategory) : (standardSimplex.obj n ⟶ X) ≃ X.obj (op n) := yonedaCompUliftFunctorEquiv X n +/-- The unique non-degenerate `n`-simplex in `Δ[n]`. -/ +def id (n : ℕ) : Δ[n] _[n] := yonedaEquiv Δ[n] [n] (𝟙 Δ[n]) + +lemma id_eq_objEquiv_symm (n : ℕ) : id n = (objEquiv _ _).symm (𝟙 _) := rfl + +lemma objEquiv_id (n : ℕ) : objEquiv _ _ (id n) = 𝟙 _ := rfl + /-- The (degenerate) `m`-simplex in the standard simplex concentrated in vertex `k`. -/ def const (n : ℕ) (k : Fin (n+1)) (m : SimplexCategoryᵒᵖ) : Δ[n].obj m := objMk (OrderHom.const _ k ) diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean index 047e74c18ff93..33e4d07da1bce 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Path.lean @@ -85,4 +85,54 @@ lemma Path.ext' {n : ℕ} {f g : Path X (n + 1)} rw [← f.arrow_tgt (Fin.last n), ← g.arrow_tgt (Fin.last n), h] · exact h j +/-- Maps of simplicial sets induce maps of paths in a simplicial set.-/ +@[simps] +def Path.map {X Y : SSet.{u}} {n : ℕ} (f : X.Path n) (σ : X ⟶ Y) : Y.Path n where + vertex i := σ.app (Opposite.op [0]) (f.vertex i) + arrow i := σ.app (Opposite.op [1]) (f.arrow i) + arrow_src i := by + simp only [← f.arrow_src i] + exact congr (σ.naturality (δ 1).op) rfl |>.symm + arrow_tgt i := by + simp only [← f.arrow_tgt i] + exact congr (σ.naturality (δ 0).op) rfl |>.symm + +/-- `Path.map` respects subintervals of paths.-/ +lemma map_interval {X Y : SSet.{u}} {n : ℕ} (f : X.Path n) (σ : X ⟶ Y) + (j l : ℕ) (hjl : j + l ≤ n) : + (f.map σ).interval j l hjl = (f.interval j l hjl).map σ := rfl + +/-- The spine of the unique non-degenerate `n`-simplex in `Δ[n]`.-/ +def standardSimplex.spineId (n : ℕ) : Path Δ[n] n := + spine Δ[n] n (standardSimplex.id n) + +/-- Any inner horn contains the spine of the unique non-degenerate `n`-simplex +in `Δ[n]`.-/ +@[simps] +def horn.spineId {n : ℕ} (i : Fin (n + 3)) + (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : + Path Λ[n + 2, i] (n + 2) where + vertex j := ⟨standardSimplex.spineId _ |>.vertex j, (horn.const n i j _).property⟩ + arrow j := ⟨standardSimplex.spineId _ |>.arrow j, by + let edge := horn.primitiveEdge h₀ hₙ j + have ha : (standardSimplex.spineId _).arrow j = edge.val := by + dsimp only [edge, standardSimplex.spineId, standardSimplex.id, spine_arrow, + mkOfSucc, horn.primitiveEdge, horn.edge, standardSimplex.edge, + standardSimplex.map_apply] + aesop + rw [ha] + exact edge.property⟩ + arrow_src := by + simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] + exact standardSimplex.spineId _ |>.arrow_src + arrow_tgt := by + simp only [horn, SimplicialObject.δ, Subtype.mk.injEq] + exact standardSimplex.spineId _ |>.arrow_tgt + +@[simp] +lemma horn.spineId_map_hornInclusion {n : ℕ} (i : Fin (n + 3)) + (h₀ : 0 < i) (hₙ : i < Fin.last (n + 2)) : + Path.map (horn.spineId i h₀ hₙ) (hornInclusion (n + 2) i) = + standardSimplex.spineId (n + 2) := rfl + end SSet diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 13a41139b88ed..ca5432b7bbe6b 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -1,10 +1,11 @@ /- Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Mario Carneiro, Emily Riehl, Joël Riou +Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path +import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Functor.KanExtension.Basic @@ -86,6 +87,147 @@ theorem spineToSimplex_edge (f : Path X n) (j l : ℕ) (hjl : j + l ≤ n) : unfold diagonal simp only [← FunctorToTypes.map_comp_apply, ← op_comp, diag_subinterval_eq] +/-- For any `σ : X ⟶ Y` between `StrictSegal` simplicial sets, `spineToSimplex` +commutes with `Path.map`. -/ +lemma spineToSimplex_map {X Y : SSet.{u}} [StrictSegal X] [StrictSegal Y] + {n : ℕ} (f : Path X (n + 1)) (σ : X ⟶ Y) : + spineToSimplex (f.map σ) = σ.app _ (spineToSimplex f) := by + apply spineInjective + ext k + dsimp only [spineEquiv, Equiv.coe_fn_mk, Path.map, spine_arrow] + rw [← types_comp_apply (σ.app _) (Y.map _), ← σ.naturality] + simp only [types_comp_apply, spineToSimplex_arrow] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +the common vertices will agree with those of the original path `f`. In particular, +a vertex `i` with `i < j` can be identified with the same vertex in `f`. -/ +lemma spine_δ_vertex_lt (f : Path X (n + 1)) {i : Fin (n + 1)} {j : Fin (n + 2)} + (h : i.castSucc < j) : + (X.spine n (X.δ j (spineToSimplex f))).vertex i = f.vertex i.castSucc := by + simp only [SimplicialObject.δ, spine_vertex] + rw [← FunctorToTypes.map_comp_apply, ← op_comp, const_comp, spineToSimplex_vertex] + simp only [SimplexCategory.δ, Hom.toOrderHom, len_mk, mkHom, Hom.mk, + OrderEmbedding.toOrderHom_coe, Fin.succAboveOrderEmb_apply] + rw [Fin.succAbove_of_castSucc_lt j i h] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +a vertex `i` with `i ≥ j` can be identified with vertex `i + 1` in the original +path. -/ +lemma spine_δ_vertex_ge (f : Path X (n + 1)) {i : Fin (n + 1)} {j : Fin (n + 2)} + (h : j ≤ i.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).vertex i = f.vertex i.succ := by + simp only [SimplicialObject.δ, spine_vertex] + rw [← FunctorToTypes.map_comp_apply, ← op_comp, const_comp, spineToSimplex_vertex] + simp only [SimplexCategory.δ, Hom.toOrderHom, len_mk, mkHom, Hom.mk, + OrderEmbedding.toOrderHom_coe, Fin.succAboveOrderEmb_apply] + rw [Fin.succAbove_of_le_castSucc j i h] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +the common arrows will agree with those of the original path `f`. In particular, +an arrow `i` with `i + 1 < j` can be identified with the same arrow in `f`. -/ +lemma spine_δ_arrow_lt (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : i.succ.castSucc < j) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = f.arrow i.castSucc := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_lt h, spineToSimplex_arrow] + +/-- If we take the path along the spine of the `j`th face of a `spineToSimplex`, +an arrow `i` with `i + 1 > j` can be identified with arrow `i + 1` in the +original path. -/ +lemma spine_δ_arrow_gt (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : j < i.succ.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = f.arrow i.succ := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_gt h, spineToSimplex_arrow] + +/-- If we take the path along the spine of a face of a `spineToSimplex`, the +arrows not contained in the original path can be recovered as the diagonal edge +of the `spineToSimplex` that "composes" arrows `i` and `i + 1`. -/ +lemma spine_δ_arrow_eq (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} + (h : j = i.succ.castSucc) : + (X.spine n (X.δ j (spineToSimplex f))).arrow i = + spineToDiagonal (Path.interval f i 2 (by omega)) := by + simp only [SimplicialObject.δ, spine_arrow] + rw [← FunctorToTypes.map_comp_apply, ← op_comp] + rw [mkOfSucc_δ_eq h, spineToSimplex_edge] + +/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ +instance : Quasicategory X := by + apply quasicategory_of_filler X + intro n i σ₀ h₀ hₙ + use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ + intro j hj + apply spineInjective + ext k + · dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + let ksucc := k.succ.castSucc + obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega + · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_lt hlt] + rfl + · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_gt hgt] + rfl + · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ + have hn0 : n ≠ 0 := by + rintro rfl + obtain rfl : k = 0 := by omega + fin_cases i <;> contradiction + /- We construct the triangle in the standard simplex as a 2-simplex in + the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, + we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ + let triangle : Λ[n + 2, i] _[2] := by + cases n with + | zero => contradiction + | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) + /- The interval spanning from `k` to `k + 2` is equivalently the spine + of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ + have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = + X.spine 2 (σ₀.app _ triangle) := by + ext m + dsimp [spine_arrow, Path.interval, Path.map] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, + whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, + Nat.reduceAdd, Quiver.Hom.unop_op] + cases n with + | zero => contradiction + | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl + rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] + simp only [spineToDiagonal, diagonal, spineToSimplex_spine] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, + uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, + len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] + ext z + cases n with + | zero => contradiction + | succ _ => + fin_cases z <;> + · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, + Quiver.Hom.unop_op, Equiv.ulift_symm_down] + rw [mkOfSucc_δ_eq heq] + rfl + end StrictSegal end SSet @@ -118,4 +260,8 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ +/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a +category is a `Quasicategory`. -/ +instance : Quasicategory (nerve C) := inferInstance + end Nerve From 8defd9fe75da7a55a73347ba859b43448550ffb5 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 11:42:18 +0000 Subject: [PATCH 137/250] chore(Filter): move defs of `Filter.IsBounded` etc (#19358) --- Mathlib/Analysis/Normed/Group/Bounded.lean | 1 - Mathlib/Order/Filter/Defs.lean | 35 ++++++++++++++++++++++ Mathlib/Order/LiminfLimsup.lean | 30 ------------------- Mathlib/Topology/Bornology/Basic.lean | 2 +- 4 files changed, 36 insertions(+), 32 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Bounded.lean b/Mathlib/Analysis/Normed/Group/Bounded.lean index dcebdaec8bfb2..0795999cb5541 100644 --- a/Mathlib/Analysis/Normed/Group/Bounded.lean +++ b/Mathlib/Analysis/Normed/Group/Bounded.lean @@ -6,7 +6,6 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import Mathlib.Analysis.Normed.Group.Basic import Mathlib.Topology.MetricSpace.Bounded import Mathlib.Order.Filter.Pointwise -import Mathlib.Order.LiminfLimsup /-! # Boundedness in normed groups diff --git a/Mathlib/Order/Filter/Defs.lean b/Mathlib/Order/Filter/Defs.lean index f03e4c29d0296..d0a00823afb8a 100644 --- a/Mathlib/Order/Filter/Defs.lean +++ b/Mathlib/Order/Filter/Defs.lean @@ -32,6 +32,11 @@ abstract two related kinds of ideas: a tactic that takes a list of proofs `hᵢ : sᵢ ∈ f`, and replaces a goal `s ∈ f` with `∀ x, x ∈ s₁ → ... → x ∈ sₙ → x ∈ s`; * `Filter.NeBot f` : a utility class stating that `f` is a non-trivial filter. +* `Filter.IsBounded r f`: the filter `f` is eventually bounded w.r.t. the relation `r`, + i.e. eventually, it is bounded by some uniform bound. + `r` will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. +* `Filter.IsCobounded r f` states that the filter `f` does not tend to infinity w.r.t. `r`. + This is also called frequently bounded. Will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. ## Notations @@ -350,6 +355,36 @@ This is essentially a push-forward along a function mapping each set to a set. - protected def lift' (f : Filter α) (h : Set α → Set β) := f.lift (𝓟 ∘ h) +/-- `f.IsBounded r`: the filter `f` is eventually bounded w.r.t. the relation `r`, +i.e. eventually, it is bounded by some uniform bound. +`r` will be usually instantiated with `(· ≤ ·)` or `(· ≥ ·)`. -/ +def IsBounded (r : α → α → Prop) (f : Filter α) := + ∃ b, ∀ᶠ x in f, r x b + +/-- `f.IsBoundedUnder (≺) u`: the image of the filter `f` under `u` is eventually bounded w.r.t. +the relation `≺`, i.e. eventually, it is bounded by some uniform bound. -/ +def IsBoundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := + (map u f).IsBounded r + +/-- `IsCobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is +also called frequently bounded. Will be usually instantiated with `≤` or `≥`. + +There is a subtlety in this definition: we want `f.IsCobounded` to hold for any `f` in the case of +complete lattices. This will be relevant to deduce theorems on complete lattices from their +versions on conditionally complete lattices with additional assumptions. We have to be careful in +the edge case of the trivial filter containing the empty set: the other natural definition + `¬ ∀ a, ∀ᶠ n in f, a ≤ n` +would not work as well in this case. +-/ +def IsCobounded (r : α → α → Prop) (f : Filter α) := + ∃ b, ∀ a, (∀ᶠ x in f, r x a) → r b a + +/-- `IsCoboundedUnder (≺) f u` states that the image of the filter `f` under the map `u` does not +tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated +with `≤` or `≥`. -/ +def IsCoboundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := + (map u f).IsCobounded r + end Filter namespace Mathlib.Tactic diff --git a/Mathlib/Order/LiminfLimsup.lean b/Mathlib/Order/LiminfLimsup.lean index 2576f34d60217..6cdea2ca4d772 100644 --- a/Mathlib/Order/LiminfLimsup.lean +++ b/Mathlib/Order/LiminfLimsup.lean @@ -46,17 +46,6 @@ namespace Filter section Relation -/-- `f.IsBounded (≺)`: the filter `f` is eventually bounded w.r.t. the relation `≺`, i.e. -eventually, it is bounded by some uniform bound. -`r` will be usually instantiated with `≤` or `≥`. -/ -def IsBounded (r : α → α → Prop) (f : Filter α) := - ∃ b, ∀ᶠ x in f, r x b - -/-- `f.IsBoundedUnder (≺) u`: the image of the filter `f` under `u` is eventually bounded w.r.t. -the relation `≺`, i.e. eventually, it is bounded by some uniform bound. -/ -def IsBoundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := - (map u f).IsBounded r - variable {r : α → α → Prop} {f g : Filter α} /-- `f` is eventually bounded if and only if, there exists an admissible set on which it is @@ -218,25 +207,6 @@ theorem IsBoundedUnder.bddBelow_range [Preorder β] [IsDirected β (· ≥ ·)] (hf : IsBoundedUnder (· ≥ ·) atTop f) : BddBelow (range f) := IsBoundedUnder.bddAbove_range (β := βᵒᵈ) hf -/-- `IsCobounded (≺) f` states that the filter `f` does not tend to infinity w.r.t. `≺`. This is -also called frequently bounded. Will be usually instantiated with `≤` or `≥`. - -There is a subtlety in this definition: we want `f.IsCobounded` to hold for any `f` in the case of -complete lattices. This will be relevant to deduce theorems on complete lattices from their -versions on conditionally complete lattices with additional assumptions. We have to be careful in -the edge case of the trivial filter containing the empty set: the other natural definition - `¬ ∀ a, ∀ᶠ n in f, a ≤ n` -would not work as well in this case. --/ -def IsCobounded (r : α → α → Prop) (f : Filter α) := - ∃ b, ∀ a, (∀ᶠ x in f, r x a) → r b a - -/-- `IsCoboundedUnder (≺) f u` states that the image of the filter `f` under the map `u` does not -tend to infinity w.r.t. `≺`. This is also called frequently bounded. Will be usually instantiated -with `≤` or `≥`. -/ -def IsCoboundedUnder (r : α → α → Prop) (f : Filter β) (u : β → α) := - (map u f).IsCobounded r - /-- To check that a filter is frequently bounded, it suffices to have a witness which bounds `f` at some point for every admissible set. diff --git a/Mathlib/Topology/Bornology/Basic.lean b/Mathlib/Topology/Bornology/Basic.lean index 6b8a44d2bdedf..e8de59a17b369 100644 --- a/Mathlib/Topology/Bornology/Basic.lean +++ b/Mathlib/Topology/Bornology/Basic.lean @@ -261,7 +261,7 @@ open Bornology theorem Filter.HasBasis.disjoint_cobounded_iff [Bornology α] {ι : Sort*} {p : ι → Prop} {s : ι → Set α} {l : Filter α} (h : l.HasBasis p s) : - Disjoint l (cobounded α) ↔ ∃ i, p i ∧ IsBounded (s i) := + Disjoint l (cobounded α) ↔ ∃ i, p i ∧ Bornology.IsBounded (s i) := h.disjoint_iff_left theorem Set.Finite.isBounded [Bornology α] {s : Set α} (hs : s.Finite) : IsBounded s := From ebd640751c72435bd1a7adfbbb1a4ee63c59cd9d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 11:42:19 +0000 Subject: [PATCH 138/250] chore: rename `Codisjoint_comm` to `codisjoint_comm` (#19403) This was a typo --- Mathlib/Data/Fintype/Basic.lean | 2 +- Mathlib/Order/BooleanAlgebra.lean | 2 +- Mathlib/Order/Disjoint.lean | 6 ++++-- Mathlib/Order/Heyting/Basic.lean | 2 +- 4 files changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index d77529bf9ac3e..d381a6d94e4b9 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -140,7 +140,7 @@ theorem codisjoint_left : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ s → a ∈ t := classical simp [codisjoint_iff, eq_univ_iff_forall, or_iff_not_imp_left] theorem codisjoint_right : Codisjoint s t ↔ ∀ ⦃a⦄, a ∉ t → a ∈ s := - Codisjoint_comm.trans codisjoint_left + codisjoint_comm.trans codisjoint_left instance booleanAlgebra [DecidableEq α] : BooleanAlgebra (Finset α) := GeneralizedBooleanAlgebra.toBooleanAlgebra diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index a2918657eefb8..cbe13b58400f7 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -681,7 +681,7 @@ theorem codisjoint_himp_self_right : Codisjoint x (x ⇨ y) := @disjoint_sdiff_self_right αᵒᵈ _ _ _ theorem himp_le : x ⇨ y ≤ z ↔ y ≤ z ∧ Codisjoint x z := - (@le_sdiff αᵒᵈ _ _ _ _).trans <| and_congr_right' <| @Codisjoint_comm _ (_) _ _ _ + (@le_sdiff αᵒᵈ _ _ _ _).trans <| and_congr_right' <| @codisjoint_comm _ (_) _ _ _ @[simp] lemma himp_le_iff : x ⇨ y ≤ x ↔ x = ⊤ := ⟨fun h ↦ codisjoint_self.1 <| codisjoint_himp_self_right.mono_right h, fun h ↦ le_top.trans h.ge⟩ diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index b024a335e595c..59bcd0815347d 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -205,12 +205,14 @@ arguments. -/ def Codisjoint (a b : α) : Prop := ∀ ⦃x⦄, a ≤ x → b ≤ x → ⊤ ≤ x -theorem Codisjoint_comm : Codisjoint a b ↔ Codisjoint b a := +theorem codisjoint_comm : Codisjoint a b ↔ Codisjoint b a := forall_congr' fun _ ↦ forall_swap +@[deprecated (since := "2024-11-23")] alias Codisjoint_comm := codisjoint_comm + @[symm] theorem Codisjoint.symm ⦃a b : α⦄ : Codisjoint a b → Codisjoint b a := - Codisjoint_comm.1 + codisjoint_comm.1 theorem symmetric_codisjoint : Symmetric (Codisjoint : α → α → Prop) := Codisjoint.symm diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index c44de8d3c3d14..1252f8884377c 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -794,7 +794,7 @@ theorem hnot_le_iff_codisjoint_right : ¬a ≤ b ↔ Codisjoint a b := by rw [← top_sdiff', sdiff_le_iff, codisjoint_iff_le_sup] theorem hnot_le_iff_codisjoint_left : ¬a ≤ b ↔ Codisjoint b a := - hnot_le_iff_codisjoint_right.trans Codisjoint_comm + hnot_le_iff_codisjoint_right.trans codisjoint_comm theorem hnot_le_comm : ¬a ≤ b ↔ ¬b ≤ a := by rw [hnot_le_iff_codisjoint_right, hnot_le_iff_codisjoint_left] From f2336c4d027db5bef5813f1a348c8df5513ad5f6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 11:42:20 +0000 Subject: [PATCH 139/250] chore: delete lemmas about T0 seminormed groups (#19438) Those duplicate lemmas about normed groups --- .../Analysis/Distribution/SchwartzSpace.lean | 2 +- .../Analysis/InnerProductSpace/Positive.lean | 2 +- .../Normed/Affine/ContinuousAffineMap.lean | 2 +- Mathlib/Analysis/Normed/Group/Basic.lean | 52 +++++++------------ Mathlib/Analysis/Normed/Group/Uniform.lean | 2 +- .../MeasureTheory/Integral/IntegrableOn.lean | 2 +- .../NumberField/CanonicalEmbedding/Basic.lean | 4 +- scripts/nolints_prime_decls.txt | 5 +- 8 files changed, 31 insertions(+), 40 deletions(-) diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 0c1edc7031b70..7aab3fe966d3b 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -1172,7 +1172,7 @@ instance instZeroAtInftyContinuousMapClass : ZeroAtInftyContinuousMapClass 𝓢( intro x hx rw [div_lt_iff₀ hε] at hx have hxpos : 0 < ‖x‖ := by - rw [norm_pos_iff'] + rw [norm_pos_iff] intro hxzero simp only [hxzero, norm_zero, zero_mul, ← not_le] at hx exact hx (apply_nonneg (SchwartzMap.seminorm ℝ 1 0) f) diff --git a/Mathlib/Analysis/InnerProductSpace/Positive.lean b/Mathlib/Analysis/InnerProductSpace/Positive.lean index f49a0c4973a85..305ed79703826 100644 --- a/Mathlib/Analysis/InnerProductSpace/Positive.lean +++ b/Mathlib/Analysis/InnerProductSpace/Positive.lean @@ -109,7 +109,7 @@ lemma antilipschitz_of_forall_le_inner_map {H : Type*} [NormedAddCommGroup H] simp_rw [sq, mul_assoc] at h by_cases hx0 : x = 0 · simp [hx0] - · apply (map_le_map_iff <| OrderIso.mulLeft₀ ‖x‖ (norm_pos_iff'.mpr hx0)).mp + · apply (map_le_map_iff <| OrderIso.mulLeft₀ ‖x‖ (norm_pos_iff.mpr hx0)).mp exact (h x).trans <| (norm_inner_le_norm _ _).trans <| (mul_comm _ _).le lemma isUnit_of_forall_le_norm_inner_map (f : E →L[𝕜] E) {c : ℝ≥0} (hc : 0 < c) diff --git a/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean index 7491177284cf8..b948db9a803a4 100644 --- a/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean +++ b/Mathlib/Analysis/Normed/Affine/ContinuousAffineMap.lean @@ -175,7 +175,7 @@ noncomputable instance : NormedAddCommGroup (V →ᴬ[𝕜] W) := simp only [norm_eq_zero, coe_const, Function.const_apply] at h₁ rw [h₁] rfl - · rw [norm_eq_zero', contLinear_eq_zero_iff_exists_const] at h₁ + · rw [norm_eq_zero, contLinear_eq_zero_iff_exists_const] at h₁ obtain ⟨q, rfl⟩ := h₁ simp only [norm_le_zero_iff, coe_const, Function.const_apply] at h₂ rw [h₂] diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index dd3bf2a5b2797..da9bbaf0e7c48 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -922,20 +922,6 @@ theorem SeminormedCommGroup.mem_closure_iff : a ∈ closure s ↔ ∀ ε, 0 < ε → ∃ b ∈ s, ‖a / b‖ < ε := by simp [Metric.mem_closure_iff, dist_eq_norm_div] -@[to_additive norm_le_zero_iff'] -theorem norm_le_zero_iff''' [T0Space E] {a : E} : ‖a‖ ≤ 0 ↔ a = 1 := by - letI : NormedGroup E := - { ‹SeminormedGroup E› with toMetricSpace := MetricSpace.ofT0PseudoMetricSpace E } - rw [← dist_one_right, dist_le_zero] - -@[to_additive norm_eq_zero'] -theorem norm_eq_zero''' [T0Space E] {a : E} : ‖a‖ = 0 ↔ a = 1 := - (norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff''' - -@[to_additive norm_pos_iff'] -theorem norm_pos_iff''' [T0Space E] {a : E} : 0 < ‖a‖ ↔ a ≠ 1 := by - rw [← not_le, norm_le_zero_iff'''] - @[to_additive] theorem SeminormedGroup.tendstoUniformlyOn_one {f : ι → κ → G} {s : Set κ} {l : Filter ι} : TendstoUniformlyOn f 1 l s ↔ ∀ ε > 0, ∀ᶠ i in l, ∀ x ∈ s, ‖f i x‖ < ε := by @@ -1292,24 +1278,26 @@ section NormedGroup variable [NormedGroup E] {a b : E} +@[to_additive (attr := simp) norm_le_zero_iff] +lemma norm_le_zero_iff' : ‖a‖ ≤ 0 ↔ a = 1 := by rw [← dist_one_right, dist_le_zero] + +@[to_additive (attr := simp) norm_pos_iff] +lemma norm_pos_iff' : 0 < ‖a‖ ↔ a ≠ 1 := by rw [← not_le, norm_le_zero_iff'] + @[to_additive (attr := simp) norm_eq_zero] -theorem norm_eq_zero'' : ‖a‖ = 0 ↔ a = 1 := - norm_eq_zero''' +lemma norm_eq_zero' : ‖a‖ = 0 ↔ a = 1 := (norm_nonneg' a).le_iff_eq.symm.trans norm_le_zero_iff' @[to_additive norm_ne_zero_iff] -theorem norm_ne_zero_iff' : ‖a‖ ≠ 0 ↔ a ≠ 1 := - norm_eq_zero''.not +lemma norm_ne_zero_iff' : ‖a‖ ≠ 0 ↔ a ≠ 1 := norm_eq_zero'.not -@[to_additive (attr := simp) norm_pos_iff] -theorem norm_pos_iff'' : 0 < ‖a‖ ↔ a ≠ 1 := - norm_pos_iff''' - -@[to_additive (attr := simp) norm_le_zero_iff] -theorem norm_le_zero_iff'' : ‖a‖ ≤ 0 ↔ a = 1 := - norm_le_zero_iff''' +@[deprecated (since := "2024-11-24")] alias norm_le_zero_iff'' := norm_le_zero_iff' +@[deprecated (since := "2024-11-24")] alias norm_le_zero_iff''' := norm_le_zero_iff' +@[deprecated (since := "2024-11-24")] alias norm_pos_iff'' := norm_pos_iff' +@[deprecated (since := "2024-11-24")] alias norm_eq_zero'' := norm_eq_zero' +@[deprecated (since := "2024-11-24")] alias norm_eq_zero''' := norm_eq_zero' @[to_additive] -theorem norm_div_eq_zero_iff : ‖a / b‖ = 0 ↔ a = b := by rw [norm_eq_zero'', div_eq_one] +theorem norm_div_eq_zero_iff : ‖a / b‖ = 0 ↔ a = b := by rw [norm_eq_zero', div_eq_one] @[to_additive] theorem norm_div_pos_iff : 0 < ‖a / b‖ ↔ a ≠ b := by @@ -1318,7 +1306,7 @@ theorem norm_div_pos_iff : 0 < ‖a / b‖ ↔ a ≠ b := by @[to_additive eq_of_norm_sub_le_zero] theorem eq_of_norm_div_le_zero (h : ‖a / b‖ ≤ 0) : a = b := by - rwa [← div_eq_one, ← norm_le_zero_iff''] + rwa [← div_eq_one, ← norm_le_zero_iff'] alias ⟨eq_of_norm_div_eq_zero, _⟩ := norm_div_eq_zero_iff @@ -1334,7 +1322,7 @@ theorem eq_one_or_nnnorm_pos (a : E) : a = 1 ∨ 0 < ‖a‖₊ := @[to_additive (attr := simp) nnnorm_eq_zero] theorem nnnorm_eq_zero' : ‖a‖₊ = 0 ↔ a = 1 := by - rw [← NNReal.coe_eq_zero, coe_nnnorm', norm_eq_zero''] + rw [← NNReal.coe_eq_zero, coe_nnnorm', norm_eq_zero'] @[to_additive nnnorm_ne_zero_iff] theorem nnnorm_ne_zero_iff' : ‖a‖₊ ≠ 0 ↔ a ≠ 1 := @@ -1346,24 +1334,24 @@ lemma nnnorm_pos' : 0 < ‖a‖₊ ↔ a ≠ 1 := pos_iff_ne_zero.trans nnnorm_n /-- See `tendsto_norm_one` for a version with full neighborhoods. -/ @[to_additive "See `tendsto_norm_zero` for a version with full neighborhoods."] lemma tendsto_norm_one' : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := - tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff''.2 hx + tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _ hx ↦ norm_pos_iff'.2 hx @[to_additive] theorem tendsto_norm_div_self_punctured_nhds (a : E) : Tendsto (fun x => ‖x / a‖) (𝓝[≠] a) (𝓝[>] 0) := (tendsto_norm_div_self a).inf <| - tendsto_principal_principal.2 fun _x hx => norm_pos_iff''.2 <| div_ne_one.2 hx + tendsto_principal_principal.2 fun _x hx => norm_pos_iff'.2 <| div_ne_one.2 hx @[to_additive] theorem tendsto_norm_nhdsWithin_one : Tendsto (norm : E → ℝ) (𝓝[≠] 1) (𝓝[>] 0) := - tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _x => norm_pos_iff''.2 + tendsto_norm_one.inf <| tendsto_principal_principal.2 fun _x => norm_pos_iff'.2 variable (E) /-- The norm of a normed group as a group norm. -/ @[to_additive "The norm of a normed group as an additive group norm."] def normGroupNorm : GroupNorm E := - { normGroupSeminorm _ with eq_one_of_map_eq_zero' := fun _ => norm_eq_zero''.1 } + { normGroupSeminorm _ with eq_one_of_map_eq_zero' := fun _ => norm_eq_zero'.1 } @[simp] theorem coe_normGroupNorm : ⇑(normGroupNorm E) = norm := diff --git a/Mathlib/Analysis/Normed/Group/Uniform.lean b/Mathlib/Analysis/Normed/Group/Uniform.lean index 73b33458e9210..ca6dd25e1c9bb 100644 --- a/Mathlib/Analysis/Normed/Group/Uniform.lean +++ b/Mathlib/Analysis/Normed/Group/Uniform.lean @@ -381,7 +381,7 @@ instance : NormedCommGroup (SeparationQuotient E) where @[to_additive] theorem mk_eq_one_iff {p : E} : mk p = 1 ↔ ‖p‖ = 0 := by - rw [← norm_mk', norm_eq_zero''] + rw [← norm_mk', norm_eq_zero'] set_option linter.docPrime false in @[to_additive (attr := simp) nnnorm_mk] diff --git a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean index f9d75e6cfea42..d23df43813b60 100644 --- a/Mathlib/MeasureTheory/Integral/IntegrableOn.lean +++ b/Mathlib/MeasureTheory/Integral/IntegrableOn.lean @@ -489,7 +489,7 @@ lemma IntegrableAtFilter.eq_zero_of_tendsto (h : IntegrableAtFilter f l μ) (h' : ∀ s ∈ l, μ s = ∞) {a : E} (hf : Tendsto f l (𝓝 a)) : a = 0 := by by_contra H - obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ε < ‖a‖ := exists_between (norm_pos_iff'.mpr H) + obtain ⟨ε, εpos, hε⟩ : ∃ (ε : ℝ), 0 < ε ∧ ε < ‖a‖ := exists_between (norm_pos_iff.mpr H) rcases h with ⟨u, ul, hu⟩ let v := u ∩ {b | ε < ‖f b‖} have hv : IntegrableOn f v μ := hu.mono_set inter_subset_left diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 3cfe928f41764..84f4e2aa14e91 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -367,8 +367,8 @@ theorem forall_normAtPlace_eq_zero_iff {x : mixedSpace K} : (∀ w, normAtPlace w x = 0) ↔ x = 0 := by refine ⟨fun h ↦ ?_, fun h ↦ ?_⟩ · ext w - · exact norm_eq_zero'.mp (normAtPlace_apply_isReal w.prop _ ▸ h w.1) - · exact norm_eq_zero'.mp (normAtPlace_apply_isComplex w.prop _ ▸ h w.1) + · exact norm_eq_zero.mp (normAtPlace_apply_isReal w.prop _ ▸ h w.1) + · exact norm_eq_zero.mp (normAtPlace_apply_isComplex w.prop _ ▸ h w.1) · simp_rw [h, map_zero, implies_true] @[deprecated (since := "2024-09-13")] alias normAtPlace_eq_zero := forall_normAtPlace_eq_zero_iff diff --git a/scripts/nolints_prime_decls.txt b/scripts/nolints_prime_decls.txt index a5ca83037daf6..414151f344621 100644 --- a/scripts/nolints_prime_decls.txt +++ b/scripts/nolints_prime_decls.txt @@ -3455,6 +3455,7 @@ NormedSpace.isVonNBounded_iff' NormedSpace.norm_expSeries_summable' NormedSpace.norm_expSeries_summable_of_mem_ball' norm_eq_of_mem_sphere' +norm_eq_zero' norm_eq_zero'' norm_eq_zero''' norm_inv' @@ -3462,6 +3463,7 @@ norm_le_norm_add_const_of_dist_le' norm_le_norm_add_norm_div' norm_le_of_mem_closedBall' norm_le_pi_norm' +norm_le_zero_iff' norm_le_zero_iff'' norm_le_zero_iff''' norm_lt_of_mem_ball' @@ -3469,7 +3471,8 @@ norm_ne_zero_iff' norm_nonneg' norm_of_subsingleton' norm_one' -norm_pos_iff'' +norm_pos_iff' +norm_pos_iff' norm_pos_iff''' norm_sub_norm_le' norm_toNNReal' From 85ded51450df1dca83cdb8736849a0028a226dcc Mon Sep 17 00:00:00 2001 From: sgouezel Date: Mon, 25 Nov 2024 11:42:22 +0000 Subject: [PATCH 140/250] =?UTF-8?q?feat:=20extend=20`ContDiff`=20to=20take?= =?UTF-8?q?=20a=20smoothness=20exponent=20in=20`WithTop=20=E2=84=95?= =?UTF-8?q?=E2=88=9E`=20(#19442)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is mathematically not meaningful at all, since the definition of `C^∞` and `C^ω` functions is the same in this PR. However, this opens the way to replacing the definition of `C^ω` with analytic in #17152. So, the current PR is just an intermediate step towards #17152, but making sure that the final diff will be smaller than the current +1280 lines changes. There is *no* mathematical content in the current PR: it's mostly adapting the library to provide an element of `WithTop ℕ∞` when it was expecting an element of `ℕ∞` before. In the locale `ContDiff`, we introduce two notations `ω` for `⊤ : WithTop ℕ∞ ` and `∞` for `(⊤ : ℕ∞) : WithTop ℕ∞`. A lot of the changes are to open the locale and use `∞` at suitable places. --- Archive/Hairer.lean | 3 +- .../Calculus/AddTorsor/AffineMap.lean | 2 +- .../Analysis/Calculus/BumpFunction/Basic.lean | 7 +- .../BumpFunction/FiniteDimension.lean | 27 +-- Mathlib/Analysis/Calculus/ContDiff/Basic.lean | 170 ++++++++++-------- .../Analysis/Calculus/ContDiff/Bounds.lean | 77 ++++---- Mathlib/Analysis/Calculus/ContDiff/Defs.lean | 114 +++++++----- .../Calculus/ContDiff/FiniteDimension.lean | 7 +- .../Analysis/Calculus/ContDiff/RCLike.lean | 8 +- .../Analysis/Calculus/FDeriv/Analytic.lean | 34 ++-- Mathlib/Analysis/Calculus/FDeriv/Norm.lean | 12 +- .../Analysis/Calculus/FDeriv/Symmetric.lean | 4 +- .../InverseFunctionTheorem/ContDiff.lean | 16 +- .../Analysis/Calculus/IteratedDeriv/Defs.lean | 22 +-- Mathlib/Analysis/Calculus/Rademacher.lean | 7 +- Mathlib/Analysis/Calculus/SmoothSeries.lean | 25 ++- Mathlib/Analysis/Calculus/Taylor.lean | 7 +- Mathlib/Analysis/Calculus/VectorField.lean | 8 +- Mathlib/Analysis/Complex/RealDeriv.lean | 4 +- Mathlib/Analysis/Convolution.lean | 9 +- .../Distribution/AEEqOfIntegralContDiff.lean | 8 +- .../Analysis/Distribution/SchwartzSpace.lean | 45 +++-- .../Fourier/FourierTransformDeriv.lean | 16 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 4 +- .../Analysis/SpecialFunctions/Log/Deriv.lean | 9 +- .../Analysis/SpecialFunctions/Pow/Deriv.lean | 7 +- .../SpecialFunctions/SmoothTransition.lean | 6 +- .../Trigonometric/ComplexDeriv.lean | 2 +- .../Trigonometric/InverseDeriv.lean | 21 ++- .../Algebra/LeftInvariantDerivation.lean | 3 + Mathlib/Geometry/Manifold/Algebra/Monoid.lean | 3 + .../Manifold/Algebra/SmoothFunctions.lean | 3 + .../Geometry/Manifold/Algebra/Structures.lean | 4 +- .../Geometry/Manifold/ContMDiff/Atlas.lean | 22 +-- .../Geometry/Manifold/ContMDiff/Basic.lean | 2 +- Mathlib/Geometry/Manifold/ContMDiff/Defs.lean | 9 +- .../Geometry/Manifold/ContMDiffMFDeriv.lean | 5 +- Mathlib/Geometry/Manifold/ContMDiffMap.lean | 4 + .../Geometry/Manifold/DerivationBundle.lean | 3 + Mathlib/Geometry/Manifold/Diffeomorph.lean | 4 +- Mathlib/Geometry/Manifold/Instances/Real.lean | 2 +- .../Geometry/Manifold/Instances/Sphere.lean | 26 +-- .../Instances/UnitsOfNormedAlgebra.lean | 3 + Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean | 6 +- Mathlib/Geometry/Manifold/MFDeriv/Basic.lean | 3 +- Mathlib/Geometry/Manifold/MFDeriv/Defs.lean | 9 +- .../Geometry/Manifold/PartitionOfUnity.lean | 3 + .../Manifold/Sheaf/LocallyRingedSpace.lean | 4 + Mathlib/Geometry/Manifold/Sheaf/Smooth.lean | 4 + .../Manifold/SmoothManifoldWithCorners.lean | 34 ++-- .../Manifold/VectorBundle/SmoothSection.lean | 3 + .../Manifold/VectorBundle/Tangent.lean | 11 +- .../Geometry/Manifold/WhitneyEmbedding.lean | 3 + MathlibTest/fun_prop2.lean | 2 +- 54 files changed, 483 insertions(+), 373 deletions(-) diff --git a/Archive/Hairer.lean b/Archive/Hairer.lean index e9b91c32132a8..6abb80f254220 100644 --- a/Archive/Hairer.lean +++ b/Archive/Hairer.lean @@ -26,6 +26,7 @@ noncomputable section open Metric Set MeasureTheory open MvPolynomial hiding support open Function hiding eval +open scoped ContDiff variable {ι : Type*} [Fintype ι] @@ -110,7 +111,7 @@ lemma inj_L : Injective (L ι) := apply subset_closure lemma hairer (N : ℕ) (ι : Type*) [Fintype ι] : - ∃ (ρ : EuclideanSpace ℝ ι → ℝ), tsupport ρ ⊆ closedBall 0 1 ∧ ContDiff ℝ ⊤ ρ ∧ + ∃ (ρ : EuclideanSpace ℝ ι → ℝ), tsupport ρ ⊆ closedBall 0 1 ∧ ContDiff ℝ ∞ ρ ∧ ∀ (p : MvPolynomial ι ℝ), p.totalDegree ≤ N → ∫ x : EuclideanSpace ℝ ι, eval x p • ρ x = eval 0 p := by have := (inj_L ι).comp (restrictTotalDegree ι ℝ N).injective_subtype diff --git a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean index 5eb7534bb2ff4..104a452fe5423 100644 --- a/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean +++ b/Mathlib/Analysis/Calculus/AddTorsor/AffineMap.lean @@ -25,7 +25,7 @@ variable [NormedAddCommGroup V] [NormedSpace 𝕜 V] variable [NormedAddCommGroup W] [NormedSpace 𝕜 W] /-- A continuous affine map between normed vector spaces is smooth. -/ -theorem contDiff {n : ℕ∞} (f : V →ᴬ[𝕜] W) : ContDiff 𝕜 n f := by +theorem contDiff {n : WithTop ℕ∞} (f : V →ᴬ[𝕜] W) : ContDiff 𝕜 n f := by rw [f.decomp] apply f.contLinear.contDiff.add exact contDiff_const diff --git a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean index 7c20ab571383f..922effc6f97d8 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/Basic.lean @@ -49,7 +49,7 @@ smooth function, smooth bump function noncomputable section open Function Set Filter -open scoped Topology Filter +open scoped Topology Filter ContDiff variable {E X : Type*} @@ -80,7 +80,7 @@ structure ContDiffBumpBase (E : Type*) [NormedAddCommGroup E] [NormedSpace ℝ E toFun : ℝ → E → ℝ mem_Icc : ∀ (R : ℝ) (x : E), toFun R x ∈ Icc (0 : ℝ) 1 symmetric : ∀ (R : ℝ) (x : E), toFun R (-x) = toFun R x - smooth : ContDiffOn ℝ ⊤ (uncurry toFun) (Ioi (1 : ℝ) ×ˢ (univ : Set E)) + smooth : ContDiffOn ℝ ∞ (uncurry toFun) (Ioi (1 : ℝ) ×ˢ (univ : Set E)) eq_one : ∀ R : ℝ, 1 < R → ∀ x : E, ‖x‖ ≤ 1 → toFun R x = 1 support : ∀ R : ℝ, 1 < R → Function.support (toFun R) = Metric.ball (0 : E) R @@ -178,7 +178,8 @@ protected theorem _root_.ContDiffWithinAt.contDiffBump {c g : X → E} {s : Set ContDiffWithinAt ℝ n (fun x => f x (g x)) s x := by change ContDiffWithinAt ℝ n (uncurry (someContDiffBumpBase E).toFun ∘ fun x : X => ((f x).rOut / (f x).rIn, (f x).rIn⁻¹ • (g x - c x))) s x - refine (((someContDiffBumpBase E).smooth.contDiffAt ?_).of_le le_top).comp_contDiffWithinAt x ?_ + refine (((someContDiffBumpBase E).smooth.contDiffAt ?_).of_le + (mod_cast le_top)).comp_contDiffWithinAt x ?_ · exact prod_mem_nhds (Ioi_mem_nhds (f x).one_lt_rOut_div_rIn) univ_mem · exact (hR.div hr (f x).rIn_pos.ne').prod ((hr.inv (f x).rIn_pos.ne').smul (hg.sub hc)) diff --git a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean index d0f390973978c..2f3e0be74bdfb 100644 --- a/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/BumpFunction/FiniteDimension.lean @@ -28,7 +28,7 @@ noncomputable section open Set Metric TopologicalSpace Function Asymptotics MeasureTheory Module ContinuousLinearMap Filter MeasureTheory.Measure Bornology -open scoped Pointwise Topology NNReal Convolution +open scoped Pointwise Topology NNReal Convolution ContDiff variable {E : Type*} [NormedAddCommGroup E] @@ -40,7 +40,7 @@ variable [NormedSpace ℝ E] [FiniteDimensional ℝ E] values in `[0, 1]`, supported in `s` and with `f x = 1`. -/ theorem exists_smooth_tsupport_subset {s : Set E} {x : E} (hs : s ∈ 𝓝 x) : ∃ f : E → ℝ, - tsupport f ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ⊤ f ∧ range f ⊆ Icc 0 1 ∧ f x = 1 := by + tsupport f ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ∞ f ∧ range f ⊆ Icc 0 1 ∧ f x = 1 := by obtain ⟨d : ℝ, d_pos : 0 < d, hd : Euclidean.closedBall x d ⊆ s⟩ := Euclidean.nhds_basis_closedBall.mem_iff.1 hs let c : ContDiffBump (toEuclidean x) := @@ -73,7 +73,7 @@ theorem exists_smooth_tsupport_subset {s : Set E} {x : E} (hs : s ∈ 𝓝 x) : /-- Given an open set `s` in a finite-dimensional real normed vector space, there exists a smooth function with values in `[0, 1]` whose support is exactly `s`. -/ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : - ∃ f : E → ℝ, f.support = s ∧ ContDiff ℝ ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := by + ∃ f : E → ℝ, f.support = s ∧ ContDiff ℝ ∞ f ∧ Set.range f ⊆ Set.Icc 0 1 := by /- For any given point `x` in `s`, one can construct a smooth function with support in `s` and nonzero at `x`. By second-countability, it follows that we may cover `s` with the supports of countably many such functions, say `g i`. @@ -85,7 +85,7 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : · exact ⟨fun _ => 0, Function.support_zero, contDiff_const, by simp only [range_const, singleton_subset_iff, left_mem_Icc, zero_le_one]⟩ - let ι := { f : E → ℝ // f.support ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ⊤ f ∧ range f ⊆ Icc 0 1 } + let ι := { f : E → ℝ // f.support ⊆ s ∧ HasCompactSupport f ∧ ContDiff ℝ ∞ f ∧ range f ⊆ Icc 0 1 } obtain ⟨T, T_count, hT⟩ : ∃ T : Set ι, T.Countable ∧ ⋃ f ∈ T, support (f : E → ℝ) = s := by have : ⋃ f : ι, (f : E → ℝ).support = s := by refine Subset.antisymm (iUnion_subset fun f => f.2.1) ?_ @@ -116,7 +116,7 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : rcases iT with ⟨n, hn⟩ rw [← hn] at hi exact ⟨n, hi⟩ - have g_smooth : ∀ n, ContDiff ℝ ⊤ (g n) := fun n => (g0 n).2.2.2.1 + have g_smooth : ∀ n, ContDiff ℝ ∞ (g n) := fun n => (g0 n).2.2.2.1 have g_comp_supp : ∀ n, HasCompactSupport (g n) := fun n => (g0 n).2.2.1 have g_nonneg : ∀ n x, 0 ≤ g n x := fun n x => ((g0 n).2.2.2.2 (mem_range_self x)).1 obtain ⟨δ, δpos, c, δc, c_lt⟩ : @@ -127,8 +127,8 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : have : ∀ i, ∃ R, ∀ x, ‖iteratedFDeriv ℝ i (fun x => g n x) x‖ ≤ R := by intro i have : BddAbove (range fun x => ‖iteratedFDeriv ℝ i (fun x : E => g n x) x‖) := by - apply - ((g_smooth n).continuous_iteratedFDeriv le_top).norm.bddAbove_range_of_hasCompactSupport + apply ((g_smooth n).continuous_iteratedFDeriv + (mod_cast le_top)).norm.bddAbove_range_of_hasCompactSupport apply HasCompactSupport.comp_left _ norm_zero apply (g_comp_supp n).iteratedFDeriv rcases this with ⟨R, hR⟩ @@ -148,7 +148,8 @@ theorem IsOpen.exists_smooth_support_eq {s : Set E} (hs : IsOpen s) : refine ⟨M⁻¹ * δ n, by positivity, fun i hi x => ?_⟩ calc ‖iteratedFDeriv ℝ i ((M⁻¹ * δ n) • g n) x‖ = ‖(M⁻¹ * δ n) • iteratedFDeriv ℝ i (g n) x‖ := by - rw [iteratedFDeriv_const_smul_apply]; exact (g_smooth n).of_le le_top + rw [iteratedFDeriv_const_smul_apply] + exact (g_smooth n).of_le (mod_cast le_top) _ = M⁻¹ * δ n * ‖iteratedFDeriv ℝ i (g n) x‖ := by rw [norm_smul _ (iteratedFDeriv ℝ i (g n) x), Real.norm_of_nonneg]; positivity _ ≤ M⁻¹ * δ n * M := (mul_le_mul_of_nonneg_left ((hR i x).trans (IR i hi)) (by positivity)) @@ -207,10 +208,10 @@ variable (E) theorem u_exists : ∃ u : E → ℝ, - ContDiff ℝ ⊤ u ∧ (∀ x, u x ∈ Icc (0 : ℝ) 1) ∧ support u = ball 0 1 ∧ ∀ x, u (-x) = u x := by + ContDiff ℝ ∞ u ∧ (∀ x, u x ∈ Icc (0 : ℝ) 1) ∧ support u = ball 0 1 ∧ ∀ x, u (-x) = u x := by have A : IsOpen (ball (0 : E) 1) := isOpen_ball obtain ⟨f, f_support, f_smooth, f_range⟩ : - ∃ f : E → ℝ, f.support = ball (0 : E) 1 ∧ ContDiff ℝ ⊤ f ∧ Set.range f ⊆ Set.Icc 0 1 := + ∃ f : E → ℝ, f.support = ball (0 : E) 1 ∧ ContDiff ℝ ∞ f ∧ Set.range f ⊆ Set.Icc 0 1 := A.exists_smooth_support_eq have B : ∀ x, f x ∈ Icc (0 : ℝ) 1 := fun x => f_range (mem_range_self x) refine ⟨fun x => (f x + f (-x)) / 2, ?_, ?_, ?_, ?_⟩ @@ -243,7 +244,7 @@ def u (x : E) : ℝ := variable (E) -theorem u_smooth : ContDiff ℝ ⊤ (u : E → ℝ) := +theorem u_smooth : ContDiff ℝ ∞ (u : E → ℝ) := (Classical.choose_spec (u_exists E)).1 theorem u_continuous : Continuous (u : E → ℝ) := @@ -433,7 +434,7 @@ theorem y_pos_of_mem_ball {D : ℝ} {x : E} (Dpos : 0 < D) (D_lt_one : D < 1) variable (E) -theorem y_smooth : ContDiffOn ℝ ⊤ (uncurry y) (Ioo (0 : ℝ) 1 ×ˢ (univ : Set E)) := by +theorem y_smooth : ContDiffOn ℝ ∞ (uncurry y) (Ioo (0 : ℝ) 1 ×ˢ (univ : Set E)) := by have hs : IsOpen (Ioo (0 : ℝ) (1 : ℝ)) := isOpen_Ioo have hk : IsCompact (closedBall (0 : E) 1) := ProperSpace.isCompact_closedBall _ _ refine contDiffOn_convolution_left_with_param (lsmul ℝ ℝ) hs hk ?_ ?_ ?_ @@ -489,7 +490,7 @@ instance (priority := 100) {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E · rfl smooth := by suffices - ContDiffOn ℝ ⊤ + ContDiffOn ℝ ∞ (uncurry y ∘ fun p : ℝ × E => ((p.1 - 1) / (p.1 + 1), ((p.1 + 1) / 2)⁻¹ • p.2)) (Ioi 1 ×ˢ univ) by apply this.congr diff --git a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean index 52ac217ea8cf2..ad3ed57edf02c 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Basic.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Basic.lean @@ -26,7 +26,7 @@ Similar results are given for `C^n` functions on domains. We use the notation `E [×n]→L[𝕜] F` for the space of continuous multilinear maps on `E^n` with values in `F`. This is the space in which the `n`-th derivative of a function from `E` to `F` lives. -In this file, we denote `⊤ : ℕ∞` with `∞`. +In this file, we denote `(⊤ : ℕ∞) : WithTop ℕ∞` with `∞` and `⊤ : WithTop ℕ∞` with `ω`. ## Tags @@ -35,9 +35,7 @@ derivative, differentiability, higher derivative, `C^n`, multilinear, Taylor ser noncomputable section -open scoped NNReal Nat - -local notation "∞" => (⊤ : ℕ∞) +open scoped NNReal Nat ContDiff universe u uE uF uG @@ -52,7 +50,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} [NormedAddCommGroup G] [NormedSpace 𝕜 G] {X : Type*} [NormedAddCommGroup X] [NormedSpace 𝕜 X] {s t : Set E} {f : E → F} - {g : F → G} {x x₀ : E} {b : E × F → G} {m n : ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} + {g : F → G} {x x₀ : E} {b : E × F → G} {m n : WithTop ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Constants -/ @@ -72,16 +70,18 @@ theorem iteratedFDeriv_zero_fun {n : ℕ} : (iteratedFDeriv 𝕜 n fun _ : E ↦ funext fun x ↦ by simpa [← iteratedFDerivWithin_univ] using iteratedFDerivWithin_zero_fun uniqueDiffOn_univ (mem_univ x) -theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := - contDiff_of_differentiable_iteratedFDeriv fun m _ => by - rw [iteratedFDeriv_zero_fun] - exact differentiable_const (0 : E[×m]→L[𝕜] F) +theorem contDiff_zero_fun : ContDiff 𝕜 n fun _ : E => (0 : F) := by + suffices ContDiff 𝕜 ω (fun _ : E => (0 : F)) from this.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega] + apply contDiff_of_differentiable_iteratedFDeriv fun m _ ↦ ?_ + rw [iteratedFDeriv_zero_fun] + exact differentiable_const (0 : E[×m]→L[𝕜] F) /-- Constants are `C^∞`. -/ theorem contDiff_const {c : F} : ContDiff 𝕜 n fun _ : E => c := by - suffices h : ContDiff 𝕜 ∞ fun _ : E => c from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω fun _ : E => c from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨differentiable_const c, ?_⟩ rw [fderiv_const] exact contDiff_zero_fun @@ -144,8 +144,8 @@ theorem contDiffWithinAt_singleton : ContDiffWithinAt 𝕜 n f {x} x := /-- Unbundled bounded linear functions are `C^∞`. -/ theorem IsBoundedLinearMap.contDiff (hf : IsBoundedLinearMap 𝕜 f) : ContDiff 𝕜 n f := by - suffices h : ContDiff 𝕜 ∞ f from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω f from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨hf.differentiable, ?_⟩ simp_rw [hf.fderiv] exact contDiff_const @@ -179,8 +179,8 @@ theorem contDiffOn_id {s} : ContDiffOn 𝕜 n (id : E → E) s := /-- Bilinear functions are `C^∞`. -/ theorem IsBoundedBilinearMap.contDiff (hb : IsBoundedBilinearMap 𝕜 b) : ContDiff 𝕜 n b := by - suffices h : ContDiff 𝕜 ∞ b from h.of_le le_top - rw [contDiff_top_iff_fderiv] + suffices h : ContDiff 𝕜 ω b from h.of_le le_top + rw [← contDiff_infty_iff_contDiff_omega, contDiff_top_iff_fderiv] refine ⟨hb.differentiable, ?_⟩ simp only [hb.fderiv] exact hb.isBoundedLinearMap_deriv.contDiff @@ -883,8 +883,8 @@ section ClmApplyConst /-- Application of a `ContinuousLinearMap` to a constant commutes with `iteratedFDerivWithin`. -/ theorem iteratedFDerivWithin_clm_apply_const_apply - {s : Set E} (hs : UniqueDiffOn 𝕜 s) {n : ℕ∞} {c : E → F →L[𝕜] G} (hc : ContDiffOn 𝕜 n c s) - {i : ℕ} (hi : i ≤ n) {x : E} (hx : x ∈ s) {u : F} {m : Fin i → E} : + {s : Set E} (hs : UniqueDiffOn 𝕜 s) {c : E → F →L[𝕜] G} + (hc : ContDiffOn 𝕜 n c s) {i : ℕ} (hi : i ≤ n) {x : E} (hx : x ∈ s) {u : F} {m : Fin i → E} : (iteratedFDerivWithin 𝕜 i (fun y ↦ (c y) u) s x) m = (iteratedFDerivWithin 𝕜 i c s x) m u := by induction i generalizing x with | zero => simp @@ -905,7 +905,7 @@ theorem iteratedFDerivWithin_clm_apply_const_apply /-- Application of a `ContinuousLinearMap` to a constant commutes with `iteratedFDeriv`. -/ theorem iteratedFDeriv_clm_apply_const_apply - {n : ℕ∞} {c : E → F →L[𝕜] G} (hc : ContDiff 𝕜 n c) + {c : E → F →L[𝕜] G} (hc : ContDiff 𝕜 n c) {i : ℕ} (hi : i ≤ n) {x : E} {u : F} {m : Fin i → E} : (iteratedFDeriv 𝕜 i (fun y ↦ (c y) u) x) m = (iteratedFDeriv 𝕜 i c x) m u := by simp only [← iteratedFDerivWithin_univ] @@ -977,7 +977,7 @@ To show that `x ↦ D_yf(x,y)g(x)` (taken within `t`) is `C^m` at `x₀` within * `g` is `C^m` at `x₀` within `s`; * Derivatives are unique at `g(x)` within `t` for `x` sufficiently close to `x₀` within `s ∪ {x₀}`; * `t` is a neighborhood of `g(x₀)` within `g '' s`; -/ -theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, UniqueDiffWithinAt 𝕜 t (g x)) (hmn : m + 1 ≤ n) @@ -990,14 +990,18 @@ theorem ContDiffWithinAt.fderivWithin'' {f : E → F → G} {g : E → F} {t : S refine hf'.congr_of_eventuallyEq_insert ?_ filter_upwards [hv, ht] exact fun y hy h2y => (hvf' y hy).fderivWithin h2y - induction' m with m - · obtain rfl := eq_top_iff.mpr hmn + match m with + | ω => + intro k hk + apply this k hk + exact le_rfl + | ∞ => rw [contDiffWithinAt_top] - exact fun m => this m le_top - exact this _ le_rfl + exact fun m => this m (mod_cast le_top) + | (m : ℕ) => exact this _ le_rfl /-- A special case of `ContDiffWithinAt.fderivWithin''` where we require that `s ⊆ g⁻¹(t)`. -/ -theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (insert x₀ s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : ∀ᶠ x in 𝓝[insert x₀ s] x₀, UniqueDiffWithinAt 𝕜 t (g x)) (hmn : m + 1 ≤ n) @@ -1006,7 +1010,7 @@ theorem ContDiffWithinAt.fderivWithin' {f : E → F → G} {g : E → F} {t : Se /-- A special case of `ContDiffWithinAt.fderivWithin'` where we require that `x₀ ∈ s` and there are unique derivatives everywhere within `t`. -/ -protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F} {t : Set F} {n : ℕ∞} +protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (ht : UniqueDiffOn 𝕜 t) (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) (hst : s ⊆ g ⁻¹' t) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 (f x) t (g x)) s x₀ := by @@ -1016,7 +1020,7 @@ protected theorem ContDiffWithinAt.fderivWithin {f : E → F → G} {g : E → F exact eventually_of_mem self_mem_nhdsWithin fun x hx => ht _ (hst hx) /-- `x ↦ fderivWithin 𝕜 (f x) t (g x) (k x)` is smooth at a point within a set. -/ -theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} {t : Set F} {n : ℕ∞} +theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} {t : Set F} (hf : ContDiffWithinAt 𝕜 n (Function.uncurry f) (s ×ˢ t) (x₀, g x₀)) (hg : ContDiffWithinAt 𝕜 m g s x₀) (hk : ContDiffWithinAt 𝕜 m k s x₀) (ht : UniqueDiffOn 𝕜 t) (hmn : m + 1 ≤ n) (hx₀ : x₀ ∈ s) (hst : s ⊆ g ⁻¹' t) : @@ -1026,7 +1030,7 @@ theorem ContDiffWithinAt.fderivWithin_apply {f : E → F → G} {g k : E → F} /-- `fderivWithin 𝕜 f s` is smooth at `x₀` within `s`. -/ theorem ContDiffWithinAt.fderivWithin_right (hf : ContDiffWithinAt 𝕜 n f s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + 1 : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (fderivWithin 𝕜 f s) s x₀ := ContDiffWithinAt.fderivWithin (ContDiffWithinAt.comp (x₀, x₀) hf contDiffWithinAt_snd <| prod_subset_preimage_snd s s) @@ -1036,7 +1040,7 @@ theorem ContDiffWithinAt.fderivWithin_right (hf : ContDiffWithinAt 𝕜 n f s x theorem ContDiffWithinAt.fderivWithin_right_apply {f : F → G} {k : F → F} {s : Set F} {n : ℕ∞} {x₀ : F} (hf : ContDiffWithinAt 𝕜 n f s x₀) (hk : ContDiffWithinAt 𝕜 m k s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + 1 : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (fun x => fderivWithin 𝕜 f s x (k x)) s x₀ := ContDiffWithinAt.fderivWithin_apply (ContDiffWithinAt.comp (x₀, x₀) hf contDiffWithinAt_snd <| prod_subset_preimage_snd s s) @@ -1044,10 +1048,10 @@ theorem ContDiffWithinAt.fderivWithin_right_apply -- TODO: can we make a version of `ContDiffWithinAt.fderivWithin` for iterated derivatives? theorem ContDiffWithinAt.iteratedFderivWithin_right {i : ℕ} (hf : ContDiffWithinAt 𝕜 n f s x₀) - (hs : UniqueDiffOn 𝕜 s) (hmn : (m + i : ℕ∞) ≤ n) (hx₀s : x₀ ∈ s) : + (hs : UniqueDiffOn 𝕜 s) (hmn : m + i ≤ n) (hx₀s : x₀ ∈ s) : ContDiffWithinAt 𝕜 m (iteratedFDerivWithin 𝕜 i f s) s x₀ := by induction' i with i hi generalizing m - · rw [ENat.coe_zero, add_zero] at hmn + · simp only [CharP.cast_eq_zero, add_zero] at hmn exact (hf.of_le hmn).continuousLinearMap_comp ((continuousMultilinearCurryFin0 𝕜 E F).symm : _ →L[𝕜] E [×0]→L[𝕜] F) · rw [Nat.cast_succ, add_comm _ 1, ← add_assoc] at hmn @@ -1056,7 +1060,7 @@ theorem ContDiffWithinAt.iteratedFderivWithin_right {i : ℕ} (hf : ContDiffWith _ →L[𝕜] E [×(i+1)]→L[𝕜] F) /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth at `x₀`. -/ -protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} +protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} (hf : ContDiffAt 𝕜 n (Function.uncurry f) (x₀, g x₀)) (hg : ContDiffAt 𝕜 m g x₀) (hmn : m + 1 ≤ n) : ContDiffAt 𝕜 m (fun x => fderiv 𝕜 (f x) (g x)) x₀ := by simp_rw [← fderivWithin_univ] @@ -1065,44 +1069,44 @@ protected theorem ContDiffAt.fderiv {f : E → F → G} {g : E → F} {n : ℕ rw [preimage_univ] /-- `fderiv 𝕜 f` is smooth at `x₀`. -/ -theorem ContDiffAt.fderiv_right (hf : ContDiffAt 𝕜 n f x₀) (hmn : (m + 1 : ℕ∞) ≤ n) : +theorem ContDiffAt.fderiv_right (hf : ContDiffAt 𝕜 n f x₀) (hmn : m + 1 ≤ n) : ContDiffAt 𝕜 m (fderiv 𝕜 f) x₀ := ContDiffAt.fderiv (ContDiffAt.comp (x₀, x₀) hf contDiffAt_snd) contDiffAt_id hmn theorem ContDiffAt.iteratedFDeriv_right {i : ℕ} (hf : ContDiffAt 𝕜 n f x₀) - (hmn : (m + i : ℕ∞) ≤ n) : ContDiffAt 𝕜 m (iteratedFDeriv 𝕜 i f) x₀ := by + (hmn : m + i ≤ n) : ContDiffAt 𝕜 m (iteratedFDeriv 𝕜 i f) x₀ := by rw [← iteratedFDerivWithin_univ, ← contDiffWithinAt_univ] at * exact hf.iteratedFderivWithin_right uniqueDiffOn_univ hmn trivial /-- `x ↦ fderiv 𝕜 (f x) (g x)` is smooth. -/ -protected theorem ContDiff.fderiv {f : E → F → G} {g : E → F} {n m : ℕ∞} +protected theorem ContDiff.fderiv {f : E → F → G} {g : E → F} (hf : ContDiff 𝕜 m <| Function.uncurry f) (hg : ContDiff 𝕜 n g) (hnm : n + 1 ≤ m) : ContDiff 𝕜 n fun x => fderiv 𝕜 (f x) (g x) := contDiff_iff_contDiffAt.mpr fun _ => hf.contDiffAt.fderiv hg.contDiffAt hnm /-- `fderiv 𝕜 f` is smooth. -/ -theorem ContDiff.fderiv_right (hf : ContDiff 𝕜 n f) (hmn : (m + 1 : ℕ∞) ≤ n) : +theorem ContDiff.fderiv_right (hf : ContDiff 𝕜 n f) (hmn : m + 1 ≤ n) : ContDiff 𝕜 m (fderiv 𝕜 f) := contDiff_iff_contDiffAt.mpr fun _x => hf.contDiffAt.fderiv_right hmn theorem ContDiff.iteratedFDeriv_right {i : ℕ} (hf : ContDiff 𝕜 n f) - (hmn : (m + i : ℕ∞) ≤ n) : ContDiff 𝕜 m (iteratedFDeriv 𝕜 i f) := + (hmn : m + i ≤ n) : ContDiff 𝕜 m (iteratedFDeriv 𝕜 i f) := contDiff_iff_contDiffAt.mpr fun _x => hf.contDiffAt.iteratedFDeriv_right hmn /-- `x ↦ fderiv 𝕜 (f x) (g x)` is continuous. -/ -theorem Continuous.fderiv {f : E → F → G} {g : E → F} {n : ℕ∞} +theorem Continuous.fderiv {f : E → F → G} {g : E → F} (hf : ContDiff 𝕜 n <| Function.uncurry f) (hg : Continuous g) (hn : 1 ≤ n) : Continuous fun x => fderiv 𝕜 (f x) (g x) := (hf.fderiv (contDiff_zero.mpr hg) hn).continuous /-- `x ↦ fderiv 𝕜 (f x) (g x) (k x)` is smooth. -/ -theorem ContDiff.fderiv_apply {f : E → F → G} {g k : E → F} {n m : ℕ∞} +theorem ContDiff.fderiv_apply {f : E → F → G} {g k : E → F} (hf : ContDiff 𝕜 m <| Function.uncurry f) (hg : ContDiff 𝕜 n g) (hk : ContDiff 𝕜 n k) (hnm : n + 1 ≤ m) : ContDiff 𝕜 n fun x => fderiv 𝕜 (f x) (g x) (k x) := (hf.fderiv hg hnm).clm_apply hk /-- The bundled derivative of a `C^{n+1}` function is `C^n`. -/ -theorem contDiffOn_fderivWithin_apply {m n : ℕ∞} {s : Set E} {f : E → F} (hf : ContDiffOn 𝕜 n f s) +theorem contDiffOn_fderivWithin_apply {s : Set E} {f : E → F} (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun p : E × E => (fderivWithin 𝕜 f s p.1 : E →L[𝕜] F) p.2) (s ×ˢ univ) := ((hf.fderivWithin hs hmn).comp contDiffOn_fst (prod_subset_preimage_fst _ _)).clm_apply @@ -1176,7 +1180,7 @@ theorem contDiffAt_pi : ContDiffAt 𝕜 n Φ x ↔ ∀ i, ContDiffAt 𝕜 n (fun theorem contDiff_pi : ContDiff 𝕜 n Φ ↔ ∀ i, ContDiff 𝕜 n fun x => Φ x i := by simp only [← contDiffOn_univ, contDiffOn_pi] -theorem contDiff_update [DecidableEq ι] (k : ℕ∞) (x : ∀ i, F' i) (i : ι) : +theorem contDiff_update [DecidableEq ι] (k : WithTop ℕ∞) (x : ∀ i, F' i) (i : ι) : ContDiff 𝕜 k (update x i) := by rw [contDiff_pi] intro j @@ -1187,7 +1191,7 @@ theorem contDiff_update [DecidableEq ι] (k : ℕ∞) (x : ∀ i, F' i) (i : ι) · exact contDiff_const variable (F') in -theorem contDiff_single [DecidableEq ι] (k : ℕ∞) (i : ι) : +theorem contDiff_single [DecidableEq ι] (k : WithTop ℕ∞) (i : ι) : ContDiff 𝕜 k (Pi.single i : F' i → ∀ i, F' i) := contDiff_update k 0 i @@ -1207,7 +1211,7 @@ section Add 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 + exact HasFTaylorSeriesUpToOn.continuousLinearMap_comp (ContinuousLinearMap.fst 𝕜 F F + .snd 𝕜 F F) (hf.prod hg) -- The sum is smooth. @@ -1637,6 +1641,10 @@ invertible element. The proof is by induction, bootstrapping using an identity derivative of inversion as a bilinear map of inversion itself. -/ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : ContDiffAt 𝕜 n Ring.inverse (x : R) := by + suffices H : ∀ (n : ℕ∞), ContDiffAt 𝕜 n Ring.inverse (x : R) by + intro k hk + exact H ⊤ k (mod_cast le_top) + intro n induction' n using ENat.nat_induction with n IH Itop · intro m hm refine ⟨{ y : R | IsUnit y }, ?_, ?_⟩ @@ -1648,7 +1656,7 @@ theorem contDiffAt_ring_inverse [CompleteSpace R] (x : Rˣ) : · rintro _ ⟨x', rfl⟩ exact (inverse_continuousAt x').continuousWithinAt · simp [ftaylorSeriesWithin] - · rw [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + · rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = 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⟩ @@ -1671,36 +1679,36 @@ variable {𝕜} -- TODO: the next few lemmas don't need `𝕜` or `𝕜'` to be complete -- A good way to show this is to generalize `contDiffAt_ring_inverse` to the setting -- of a function `f` such that `∀ᶠ x in 𝓝 a, x * f x = 1`. -theorem ContDiffWithinAt.inv {f : E → 𝕜'} {n} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : +theorem ContDiffWithinAt.inv {f : E → 𝕜'} (hf : ContDiffWithinAt 𝕜 n f s x) (hx : f x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => (f x)⁻¹) s x := (contDiffAt_inv 𝕜 hx).comp_contDiffWithinAt x hf -theorem ContDiffOn.inv {f : E → 𝕜'} {n} (hf : ContDiffOn 𝕜 n f s) (h : ∀ x ∈ s, f x ≠ 0) : +theorem ContDiffOn.inv {f : E → 𝕜'} (hf : ContDiffOn 𝕜 n f s) (h : ∀ x ∈ s, f x ≠ 0) : ContDiffOn 𝕜 n (fun x => (f x)⁻¹) s := fun x hx => (hf.contDiffWithinAt hx).inv (h x hx) -nonrec theorem ContDiffAt.inv {f : E → 𝕜'} {n} (hf : ContDiffAt 𝕜 n f x) (hx : f x ≠ 0) : +nonrec theorem ContDiffAt.inv {f : E → 𝕜'} (hf : ContDiffAt 𝕜 n f x) (hx : f x ≠ 0) : ContDiffAt 𝕜 n (fun x => (f x)⁻¹) x := hf.inv hx -theorem ContDiff.inv {f : E → 𝕜'} {n} (hf : ContDiff 𝕜 n f) (h : ∀ x, f x ≠ 0) : +theorem ContDiff.inv {f : E → 𝕜'} (hf : ContDiff 𝕜 n f) (h : ∀ x, f x ≠ 0) : ContDiff 𝕜 n fun x => (f x)⁻¹ := by rw [contDiff_iff_contDiffAt]; exact fun x => hf.contDiffAt.inv (h x) -- TODO: generalize to `f g : E → 𝕜'` -theorem ContDiffWithinAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffWithinAt 𝕜 n f s x) +theorem ContDiffWithinAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffWithinAt 𝕜 n f s x) (hg : ContDiffWithinAt 𝕜 n g s x) (hx : g x ≠ 0) : ContDiffWithinAt 𝕜 n (fun x => f x / g x) s x := by simpa only [div_eq_mul_inv] using hf.mul (hg.inv hx) -theorem ContDiffOn.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffOn 𝕜 n f s) +theorem ContDiffOn.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffOn 𝕜 n f s) (hg : ContDiffOn 𝕜 n g s) (h₀ : ∀ x ∈ s, g x ≠ 0) : ContDiffOn 𝕜 n (f / g) s := fun x hx => (hf x hx).div (hg x hx) (h₀ x hx) -nonrec theorem ContDiffAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiffAt 𝕜 n f x) +nonrec theorem ContDiffAt.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiffAt 𝕜 n f x) (hg : ContDiffAt 𝕜 n g x) (hx : g x ≠ 0) : ContDiffAt 𝕜 n (fun x => f x / g x) x := hf.div hg hx -theorem ContDiff.div [CompleteSpace 𝕜] {f g : E → 𝕜} {n} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) +theorem ContDiff.div [CompleteSpace 𝕜] {f g : E → 𝕜} (hf : ContDiff 𝕜 n f) (hg : ContDiff 𝕜 n g) (h0 : ∀ x, g x ≠ 0) : ContDiff 𝕜 n fun x => f x / g x := by simp only [contDiff_iff_contDiffAt] at * exact fun x => (hf x).div (hg x) (h0 x) @@ -1738,23 +1746,17 @@ section FunctionInverse open ContinuousLinearMap -/-- If `f` is a local homeomorphism and the point `a` is in its target, -and if `f` is `n` times continuously differentiable at `f.symm a`, -and if the derivative at `f.symm a` is a continuous linear equivalence, -then `f.symm` is `n` times continuously differentiable at the point `a`. - -This is one of the easy parts of the inverse function theorem: it assumes that we already have -an inverse function. -/ -theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomorph E F) +private theorem PartialHomeomorph.contDiffAt_symm_aux {n : ℕ∞} + [CompleteSpace E] (f : PartialHomeomorph E F) {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : ContDiffAt 𝕜 n f.symm a := by - -- We prove this by induction on `n` + -- We prove this by induction on `n` induction' n using ENat.nat_induction with n IH Itop - · rw [contDiffAt_zero] + · apply contDiffAt_zero.2 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 [show (n.succ : ℕ∞) = n + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + rw [show ((n.succ : ℕ∞) : WithTop ℕ∞) = 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₀' @@ -1791,6 +1793,24 @@ theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomor intro n exact Itop n (contDiffAt_top.mp hf n) +/-- If `f` is a local homeomorphism and the point `a` is in its target, +and if `f` is `n` times continuously differentiable at `f.symm a`, +and if the derivative at `f.symm a` is a continuous linear equivalence, +then `f.symm` is `n` times continuously differentiable at the point `a`. + +This is one of the easy parts of the inverse function theorem: it assumes that we already have +an inverse function. -/ +theorem PartialHomeomorph.contDiffAt_symm [CompleteSpace E] (f : PartialHomeomorph E F) + {f₀' : E ≃L[𝕜] F} {a : F} (ha : a ∈ f.target) + (hf₀' : HasFDerivAt f (f₀' : E →L[𝕜] F) (f.symm a)) (hf : ContDiffAt 𝕜 n f (f.symm a)) : + ContDiffAt 𝕜 n f.symm a := by + match n with + | ω => + intro k hk + exact f.contDiffAt_symm_aux ha hf₀' (hf.of_le (m := k) le_top) k le_rfl + | (n : ℕ∞) => + exact f.contDiffAt_symm_aux ha hf₀' hf + /-- If `f` is an `n` times continuously differentiable homeomorphism, and if the derivative of `f` at each point is a continuous linear equivalence, then `f.symm` is `n` times continuously differentiable. @@ -1905,14 +1925,14 @@ theorem contDiffOn_top_iff_derivWithin (hs : UniqueDiffOn 𝕜 s₂) : ContDiffOn 𝕜 ∞ f₂ s₂ ↔ DifferentiableOn 𝕜 f₂ s₂ ∧ ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := by constructor · intro h - refine ⟨h.differentiableOn le_top, ?_⟩ + refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_derivWithin hs).1 ?_).2 - exact h.of_le le_top + exact h.of_le (mod_cast le_top) · intro h refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := le_top + have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top apply ((contDiffOn_succ_iff_derivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact WithTop.coe_le_coe.2 (Nat.le_succ n) + exact_mod_cast (Nat.le_succ n) /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (formulated with `deriv`) is `C^∞`. -/ @@ -1923,13 +1943,17 @@ theorem contDiffOn_top_iff_deriv_of_isOpen (hs : IsOpen s₂) : protected theorem ContDiffOn.derivWithin (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : UniqueDiffOn 𝕜 s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (derivWithin f₂ s₂) s₂ := by - cases m - · change ∞ + 1 ≤ n at hmn - have : n = ∞ := by simpa using hmn - rw [this] at hf - exact ((contDiffOn_top_iff_derivWithin hs).1 hf).2 - · change (Nat.succ _ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 + rcases le_or_lt ∞ n with hn | hn + · have : ContDiffOn 𝕜 ∞ (derivWithin f₂ s₂) s₂ := + ((contDiffOn_top_iff_derivWithin hs).1 (hf.of_le hn)).2 + intro x hx k hk + exact this x hx k (mod_cast le_top) + · match m with + | ω => simpa using hmn.trans_lt hn + | ∞ => simpa using hmn.trans_lt hn + | (m : ℕ) => + change (m.succ : ℕ∞) ≤ n at hmn + exact ((contDiffOn_succ_iff_derivWithin hs).1 (hf.of_le hmn)).2 theorem ContDiffOn.deriv_of_isOpen (hf : ContDiffOn 𝕜 n f₂ s₂) (hs : IsOpen s₂) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (deriv f₂) s₂ := diff --git a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean index 3d7d145291bee..fe52bff69b40d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Bounds.lean @@ -53,7 +53,7 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu · simp only [norm_iteratedFDerivWithin_zero, zero_add, Finset.range_one, Finset.sum_singleton, Nat.choose_self, Nat.cast_one, one_mul, Nat.sub_zero, ← mul_assoc] apply B.le_opNorm₂ - · have In : (n : ℕ∞) + 1 ≤ n.succ := by simp only [Nat.cast_succ, le_refl] + · have In : (n : WithTop ℕ∞) + 1 ≤ n.succ := by simp only [Nat.cast_succ, le_refl] -- Porting note: the next line is a hack allowing Lean to find the operator norm instance. let norm := @ContinuousLinearMap.hasOpNorm _ _ Eu ((Du →L[𝕜] Fu) →L[𝕜] Du →L[𝕜] Gu) _ _ _ _ _ _ (RingHom.id 𝕜) @@ -101,7 +101,7 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu iteratedFDerivWithin 𝕜 n (fun y => B.precompR Du (f y) (fderivWithin 𝕜 g s y) + B.precompL Du (fderivWithin 𝕜 f s y) (g y)) s x := by apply iteratedFDerivWithin_congr (fun y hy => ?_) hx - have L : (1 : ℕ∞) ≤ n.succ := by + have L : (1 : WithTop ℕ∞) ≤ n.succ := by simpa only [ENat.coe_one, Nat.one_le_cast] using Nat.succ_pos n exact B.fderivWithin_of_bilinear (hf.differentiableOn L y hy) (hg.differentiableOn L y hy) (hs y hy) @@ -123,8 +123,8 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_aux {Du Eu iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L[𝕜] F →L[𝕜] G) - {f : D → E} {g : D → F} {N : ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) - (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : + {f : D → E} {g : D → F} {N : WithTop ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) + (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ ‖B‖ * ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by @@ -205,8 +205,8 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear (B : E →L iterated derivatives of `f` and `g` when `B` is bilinear: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ‖B‖ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} - {g : D → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : + {g : D → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} + (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ ‖B‖ * ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -217,9 +217,9 @@ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear (B : E →L[𝕜] iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one - (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : ℕ∞} {s : Set D} {x : D} + (B : E →L[𝕜] F →L[𝕜] G) {f : D → E} {g : D → F} {N : WithTop ℕ∞} {s : Set D} {x : D} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ + (hn : n ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDerivWithin 𝕜 n (fun y => B (f y) (g y)) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by apply (B.norm_iteratedFDerivWithin_le_of_bilinear hf hg hs hx hn).trans @@ -229,8 +229,9 @@ theorem ContinuousLinearMap.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one iterated derivatives of `f` and `g` when `B` is bilinear of norm at most `1`: `‖D^n (x ↦ B (f x) (g x))‖ ≤ ∑_{k ≤ n} n.choose k ‖D^k f‖ ‖D^{n-k} g‖` -/ theorem ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear_of_le_one (B : E →L[𝕜] F →L[𝕜] G) - {f : D → E} {g : D → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : D) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) (hB : ‖B‖ ≤ 1) : ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ + {f : D → E} {g : D → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) + (x : D) {n : ℕ} (hn : n ≤ N) (hB : ‖B‖ ≤ 1) : + ‖iteratedFDeriv 𝕜 n (fun y => B (f y) (g y)) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -242,17 +243,17 @@ section variable {𝕜' : Type*} [NormedField 𝕜'] [NormedAlgebra 𝕜 𝕜'] [NormedSpace 𝕜' F] [IsScalarTower 𝕜 𝕜' F] -theorem norm_iteratedFDerivWithin_smul_le {f : E → 𝕜'} {g : E → F} {N : ℕ∞} +theorem norm_iteratedFDerivWithin_smul_le {f : E → 𝕜'} {g : E → F} {N : WithTop ℕ∞} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) - {n : ℕ} (hn : (n : ℕ∞) ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y • g y) s x‖ ≤ + {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y • g y) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := (ContinuousLinearMap.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F).norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn ContinuousLinearMap.opNorm_lsmul_le -theorem norm_iteratedFDeriv_smul_le {f : E → 𝕜'} {g : E → F} {N : ℕ∞} (hf : ContDiff 𝕜 N f) - (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : +theorem norm_iteratedFDeriv_smul_le {f : E → 𝕜'} {g : E → F} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) + (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => f y • g y) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := (ContinuousLinearMap.lsmul 𝕜 𝕜' : 𝕜' →L[𝕜] F →L[𝕜] F).norm_iteratedFDeriv_le_of_bilinear_of_le_one @@ -265,17 +266,18 @@ section variable {ι : Type*} {A : Type*} [NormedRing A] [NormedAlgebra 𝕜 A] {A' : Type*} [NormedCommRing A'] [NormedAlgebra 𝕜 A'] -theorem norm_iteratedFDerivWithin_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf : ContDiffOn 𝕜 N f s) - (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => f y * g y) s x‖ ≤ +theorem norm_iteratedFDerivWithin_mul_le {f : E → A} {g : E → A} {N : WithTop ℕ∞} + (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) + {x : E} (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : + ‖iteratedFDerivWithin 𝕜 n (fun y => f y * g y) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := (ContinuousLinearMap.mul 𝕜 A : A →L[𝕜] A →L[𝕜] A).norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn (ContinuousLinearMap.opNorm_mul_le _ _) -theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf : ContDiff 𝕜 N f) - (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : +theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : WithTop ℕ∞} (hf : ContDiff 𝕜 N f) + (hg : ContDiff 𝕜 N g) (x : E) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y => f y * g y) x‖ ≤ ∑ i ∈ Finset.range (n + 1), (n.choose i : ℝ) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp_rw [← iteratedFDerivWithin_univ] @@ -285,8 +287,8 @@ theorem norm_iteratedFDeriv_mul_le {f : E → A} {g : E → A} {N : ℕ∞} (hf -- TODO: Add `norm_iteratedFDeriv[Within]_list_prod_le` for non-commutative `NormedRing A`. theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u : Finset ι} - {f : ι → E → A'} {N : ℕ∞} (hf : ∀ i ∈ u, ContDiffOn 𝕜 N (f i) s) (hs : UniqueDiffOn 𝕜 s) {x : E} - (hx : x ∈ s) {n : ℕ} (hn : (n : ℕ∞) ≤ N) : + {f : ι → E → A'} {N : WithTop ℕ∞} (hf : ∀ i ∈ u, ContDiffOn 𝕜 N (f i) s) + (hs : UniqueDiffOn 𝕜 s) {x : E} (hx : x ∈ s) {n : ℕ} (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (∏ j ∈ u, f j ·) s x‖ ≤ ∑ p ∈ u.sym n, (p : Multiset ι).multinomial * ∏ j ∈ u, ‖iteratedFDerivWithin 𝕜 (Multiset.count j p) (f j) s x‖ := by @@ -309,7 +311,7 @@ theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u simp only [comp_apply, Finset.symInsertEquiv_symm_apply_coe] refine Finset.sum_le_sum ?_ intro m _ - specialize IH hf.2 (n := n - m) (le_trans (WithTop.coe_le_coe.mpr (n.sub_le m)) hn) + specialize IH hf.2 (n := n - m) (le_trans (by exact_mod_cast n.sub_le m) hn) refine le_trans (mul_le_mul_of_nonneg_left IH (by simp [mul_nonneg])) ?_ rw [Finset.mul_sum, ← Finset.sum_coe_sort] refine Finset.sum_le_sum ?_ @@ -329,8 +331,8 @@ theorem norm_iteratedFDerivWithin_prod_le [DecidableEq ι] [NormOneClass A'] {u rw [Sym.count_coe_fill_of_ne hji] theorem norm_iteratedFDeriv_prod_le [DecidableEq ι] [NormOneClass A'] {u : Finset ι} - {f : ι → E → A'} {N : ℕ∞} (hf : ∀ i ∈ u, ContDiff 𝕜 N (f i)) {x : E} {n : ℕ} - (hn : (n : ℕ∞) ≤ N) : + {f : ι → E → A'} {N : WithTop ℕ∞} (hf : ∀ i ∈ u, ContDiff 𝕜 N (f i)) {x : E} {n : ℕ} + (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (∏ j ∈ u, f j ·) x‖ ≤ ∑ p ∈ u.sym n, (p : Multiset ι).multinomial * ∏ j ∈ u, ‖iteratedFDeriv 𝕜 ((p : Multiset ι).count j) (f j) x‖ := by @@ -362,7 +364,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro induction' n using Nat.case_strong_induction_on with n IH generalizing Gu · simpa [norm_iteratedFDerivWithin_zero, Nat.factorial_zero, algebraMap.coe_one, one_mul, pow_zero, mul_one, comp_apply] using hC 0 le_rfl - have M : (n : ℕ∞) < n.succ := Nat.cast_lt.2 n.lt_succ_self + have M : (n : WithTop ℕ∞) < n.succ := Nat.cast_lt.2 n.lt_succ_self have Cnonneg : 0 ≤ C := (norm_nonneg _).trans (hC 0 bot_le) have Dnonneg : 0 ≤ D := by have : 1 ≤ n + 1 := by simp only [le_add_iff_nonneg_left, zero_le'] @@ -404,7 +406,8 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro LinearIsometryEquiv.norm_map] _ = ‖iteratedFDerivWithin 𝕜 n (fun y : E => ContinuousLinearMap.compL 𝕜 E Fu Gu (fderivWithin 𝕜 g t (f y)) (fderivWithin 𝕜 f s y)) s x‖ := by - have L : (1 : ℕ∞) ≤ n.succ := by simpa only [ENat.coe_one, Nat.one_le_cast] using n.succ_pos + have L : (1 : WithTop ℕ∞) ≤ n.succ := by + simpa only [ENat.coe_one, Nat.one_le_cast] using n.succ_pos congr 1 refine iteratedFDerivWithin_congr (fun y hy => ?_) hx _ apply fderivWithin_comp _ _ _ hst (hs y hy) @@ -459,7 +462,7 @@ theorem norm_iteratedFDerivWithin_comp_le_aux {Fu Gu : Type u} [NormedAddCommGro within a set of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative of `g ∘ f` is bounded by `n! * C * D^n`. -/ theorem norm_iteratedFDerivWithin_comp_le {g : F → G} {f : E → F} {n : ℕ} {s : Set E} {t : Set F} - {x : E} {N : ℕ∞} (hg : ContDiffOn 𝕜 N g t) (hf : ContDiffOn 𝕜 N f s) (hn : (n : ℕ∞) ≤ N) + {x : E} {N : WithTop ℕ∞} (hg : ContDiffOn 𝕜 N g t) (hf : ContDiffOn 𝕜 N f s) (hn : n ≤ N) (ht : UniqueDiffOn 𝕜 t) (hs : UniqueDiffOn 𝕜 s) (hst : MapsTo f s t) (hx : x ∈ s) {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iteratedFDerivWithin 𝕜 i g t (f x)‖ ≤ C) (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iteratedFDerivWithin 𝕜 i f s x‖ ≤ D ^ i) : @@ -509,8 +512,8 @@ theorem norm_iteratedFDerivWithin_comp_le {g : F → G} {f : E → F} {n : ℕ} /-- If the derivatives of `g` at `f x` are bounded by `C`, and the `i`-th derivative of `f` at `x` is bounded by `D^i` for all `1 ≤ i ≤ n`, then the `n`-th derivative of `g ∘ f` is bounded by `n! * C * D^n`. -/ -theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : ℕ∞} (hg : ContDiff 𝕜 N g) - (hf : ContDiff 𝕜 N f) (hn : (n : ℕ∞) ≤ N) (x : E) {C : ℝ} {D : ℝ} +theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : WithTop ℕ∞} + (hg : ContDiff 𝕜 N g) (hf : ContDiff 𝕜 N f) (hn : n ≤ N) (x : E) {C : ℝ} {D : ℝ} (hC : ∀ i, i ≤ n → ‖iteratedFDeriv 𝕜 i g (f x)‖ ≤ C) (hD : ∀ i, 1 ≤ i → i ≤ n → ‖iteratedFDeriv 𝕜 i f x‖ ≤ D ^ i) : ‖iteratedFDeriv 𝕜 n (g ∘ f) x‖ ≤ n ! * C * D ^ n := by @@ -521,8 +524,9 @@ theorem norm_iteratedFDeriv_comp_le {g : F → G} {f : E → F} {n : ℕ} {N : section Apply theorem norm_iteratedFDerivWithin_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {s : Set E} {x : E} - {N : ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) (hs : UniqueDiffOn 𝕜 s) - (hx : x ∈ s) (hn : ↑n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y => (f y) (g y)) s x‖ ≤ + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hg : ContDiffOn 𝕜 N g s) + (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hn : n ≤ N) : + ‖iteratedFDerivWithin 𝕜 n (fun y => (f y) (g y)) s x‖ ≤ ∑ i ∈ Finset.range (n + 1), ↑(n.choose i) * ‖iteratedFDerivWithin 𝕜 i f s x‖ * ‖iteratedFDerivWithin 𝕜 (n - i) g s x‖ := by let B : (F →L[𝕜] G) →L[𝕜] F →L[𝕜] G := ContinuousLinearMap.flip (ContinuousLinearMap.apply 𝕜 G) @@ -533,8 +537,8 @@ theorem norm_iteratedFDerivWithin_clm_apply {f : E → F →L[𝕜] G} {g : E rfl exact B.norm_iteratedFDerivWithin_le_of_bilinear_of_le_one hf hg hs hx hn hB -theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {N : ℕ∞} {n : ℕ} - (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : E) (hn : ↑n ≤ N) : +theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} {N : WithTop ℕ∞} {n : ℕ} + (hf : ContDiff 𝕜 N f) (hg : ContDiff 𝕜 N g) (x : E) (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y : E => (f y) (g y)) x‖ ≤ ∑ i ∈ Finset.range (n + 1), ↑(n.choose i) * ‖iteratedFDeriv 𝕜 i f x‖ * ‖iteratedFDeriv 𝕜 (n - i) g x‖ := by simp only [← iteratedFDerivWithin_univ] @@ -542,7 +546,8 @@ theorem norm_iteratedFDeriv_clm_apply {f : E → F →L[𝕜] G} {g : E → F} { (Set.mem_univ x) hn theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {s : Set E} {x : E} - {N : ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ s) (hn : ↑n ≤ N) : + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiffOn 𝕜 N f s) (hs : UniqueDiffOn 𝕜 s) + (hx : x ∈ s) (hn : n ≤ N) : ‖iteratedFDerivWithin 𝕜 n (fun y : E => (f y) c) s x‖ ≤ ‖c‖ * ‖iteratedFDerivWithin 𝕜 n f s x‖ := by let g : (F →L[𝕜] G) →L[𝕜] G := ContinuousLinearMap.apply 𝕜 G c @@ -553,8 +558,8 @@ theorem norm_iteratedFDerivWithin_clm_apply_const {f : E → F →L[𝕜] G} {c rw [ContinuousLinearMap.apply_apply, mul_comm] exact f.le_opNorm c -theorem norm_iteratedFDeriv_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {x : E} {N : ℕ∞} {n : ℕ} - (hf : ContDiff 𝕜 N f) (hn : ↑n ≤ N) : +theorem norm_iteratedFDeriv_clm_apply_const {f : E → F →L[𝕜] G} {c : F} {x : E} + {N : WithTop ℕ∞} {n : ℕ} (hf : ContDiff 𝕜 N f) (hn : n ≤ N) : ‖iteratedFDeriv 𝕜 n (fun y : E => (f y) c) x‖ ≤ ‖c‖ * ‖iteratedFDeriv 𝕜 n f x‖ := by simp only [← iteratedFDerivWithin_univ] exact norm_iteratedFDerivWithin_clm_apply_const hf.contDiffOn uniqueDiffOn_univ diff --git a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean index 968b9b7d7d1c3..180ea22be1a5e 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/Defs.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/Defs.lean @@ -93,7 +93,13 @@ noncomputable section open scoped Classical open NNReal Topology Filter -local notation "∞" => (⊤ : ℕ∞) +/-- Smoothness exponent for analytic functions. Not implemented yet, `ω` smoothness is equivalent +to `∞` smoothness in current mathlib. -/ +scoped [ContDiff] notation3 "ω" => (⊤ : WithTop ℕ∞) +/-- Smoothness exponent for infinitely differentiable functions. -/ +scoped [ContDiff] notation3 "∞" => ((⊤ : ℕ∞) : WithTop ℕ∞) + +open ContDiff /- Porting note: These lines are not required in Mathlib4. @@ -108,7 +114,7 @@ universe u uE uF uG uX variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] {G : Type uG} [NormedAddCommGroup G] [NormedSpace 𝕜 G] {X : Type uX} [NormedAddCommGroup X] [NormedSpace 𝕜 X] - {s s₁ t u : Set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {m n : ℕ∞} + {s s₁ t u : Set E} {f f₁ : E → F} {g : F → G} {x x₀ : E} {c : F} {m n : WithTop ℕ∞} {p : E → FormalMultilinearSeries 𝕜 E F} /-! ### Smooth functions within a set around a point -/ @@ -122,8 +128,11 @@ depend on the finite order we consider). For instance, a real function which is `C^m` on `(-1/m, 1/m)` for each natural `m`, but not better, is `C^∞` at `0` within `univ`. + +We take the exponent `n` in `WithTop ℕ∞` to allow for an extension to analytic functions in the +future, but currently the notion is the same for `n = ∞` and `n = ω`. -/ -def ContDiffWithinAt (n : ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := +def ContDiffWithinAt (n : WithTop ℕ∞) (f : E → F) (s : Set E) (x : E) : Prop := ∀ m : ℕ, m ≤ n → ∃ u ∈ 𝓝[insert x s] x, ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpToOn m f p u @@ -137,9 +146,9 @@ theorem contDiffWithinAt_nat {n : ℕ} : 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) -theorem contDiffWithinAt_iff_forall_nat_le : +theorem contDiffWithinAt_iff_forall_nat_le {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffWithinAt 𝕜 m f s x := - ⟨fun H _m hm => H.of_le hm, fun H m hm => H m hm _ le_rfl⟩ + ⟨fun H _m hm => H.of_le (mod_cast hm), fun H m hm => H m (mod_cast hm) _ le_rfl⟩ theorem contDiffWithinAt_top : ContDiffWithinAt 𝕜 ∞ f s x ↔ ∀ n : ℕ, ContDiffWithinAt 𝕜 n f s x := contDiffWithinAt_iff_forall_nat_le.trans <| by simp only [forall_prop_of_true, le_top] @@ -308,7 +317,7 @@ theorem contDiffWithinAt_succ_iff_hasFDerivWithinAt {n : ℕ} : hasFTaylorSeriesUpToOn_succ_iff_right] at Hp exact Hp.2.2.of_le (mod_cast hm) · rintro ⟨u, hu, f', f'_eq_deriv, Hf'⟩ - rw [show (n : ℕ∞) + 1 = (n + 1 : ℕ) from rfl, contDiffWithinAt_nat] + rw [show (n : WithTop ℕ∞) + 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 @@ -370,7 +379,7 @@ admits continuous derivatives up to order `n` on a neighborhood of `x` in `s`. For `n = ∞`, we only require that this holds up to any finite order (where the neighborhood may depend on the finite order we consider). -/ -def ContDiffOn (n : ℕ∞) (f : E → F) (s : Set E) : Prop := +def ContDiffOn (n : WithTop ℕ∞) (f : E → F) (s : Set E) : Prop := ∀ x ∈ s, ContDiffWithinAt 𝕜 n f s x variable {𝕜} @@ -427,13 +436,15 @@ theorem ContDiffOn.of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDif theorem ContDiffOn.one_of_succ {n : ℕ} (h : ContDiffOn 𝕜 (n + 1) f s) : ContDiffOn 𝕜 1 f s := h.of_le <| WithTop.coe_le_coe.mpr le_add_self -theorem contDiffOn_iff_forall_nat_le : ContDiffOn 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffOn 𝕜 m f s := - ⟨fun H _ hm => H.of_le hm, fun H x hx m hm => H m hm x hx m le_rfl⟩ +theorem contDiffOn_iff_forall_nat_le {n : ℕ∞} : + ContDiffOn 𝕜 n f s ↔ ∀ m : ℕ, ↑m ≤ n → ContDiffOn 𝕜 m f s := + ⟨fun H _ hm => H.of_le (mod_cast hm), fun H x hx m hm => H m (mod_cast hm) x hx m le_rfl⟩ theorem contDiffOn_top : ContDiffOn 𝕜 ∞ f s ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := contDiffOn_iff_forall_nat_le.trans <| by simp only [le_top, forall_prop_of_true] -theorem contDiffOn_all_iff_nat : (∀ n, ContDiffOn 𝕜 n f s) ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := by +theorem contDiffOn_all_iff_nat : + (∀ (n : ℕ∞), ContDiffOn 𝕜 n f s) ↔ ∀ n : ℕ, ContDiffOn 𝕜 n f s := by refine ⟨fun H n => H n, ?_⟩ rintro H (_ | n) exacts [contDiffOn_top.2 H, H n] @@ -522,8 +533,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 - have : m < n := mod_cast hm - rcases (h x hx) m.succ (Order.add_one_le_of_lt this) with ⟨u, hu, p, Hp⟩ + have : (m + 1 : ℕ) ≤ n := ENat.add_one_natCast_le_withTop_of_lt hm + rcases (h x hx).of_le this _ le_rfl 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 @@ -557,7 +568,7 @@ protected theorem ContDiffOn.ftaylorSeriesWithin (h : ContDiffOn 𝕜 n f s) (hs exact (Hp.mono ho).eq_iteratedFDerivWithin_of_uniqueDiffOn le_rfl (hs.inter o_open) ⟨hy, yo⟩ exact ((Hp.mono ho).cont m le_rfl).congr fun y hy => (A y hy).symm -theorem contDiffOn_of_continuousOn_differentiableOn +theorem contDiffOn_of_continuousOn_differentiableOn {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) : @@ -570,11 +581,11 @@ 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 (mod_cast hk) hm) y hy).hasFDerivWithinAt + convert (Hdiff k (lt_of_lt_of_le (mod_cast hk) (mod_cast hm)) y hy).hasFDerivWithinAt · intro k hk - exact Hcont k (le_trans (mod_cast hk) hm) + exact Hcont k (le_trans (mod_cast hk) (mod_cast hm)) -theorem contDiffOn_of_differentiableOn +theorem contDiffOn_of_differentiableOn {n : ℕ∞} (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 => @@ -592,7 +603,7 @@ theorem ContDiffOn.differentiableOn_iteratedFDerivWithin {m : ℕ} (h : ContDiff theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} (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⟩ + rcases h.contDiffOn' (ENat.add_one_natCast_le_withTop_of_lt hmn) with ⟨u, uo, xu, hu⟩ set t := insert x s ∩ u have A : t =ᶠ[𝓝[≠] x] s := by simp only [set_eventuallyEq_iff_inf_principal, ← nhdsWithin_inter'] @@ -607,12 +618,12 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedFDerivWithin {m : ℕ} rw [differentiableWithinAt_congr_set' _ A] at C exact C.congr_of_eventuallyEq (B.filter_mono inf_le_left) B.self_of_nhds -theorem contDiffOn_iff_continuousOn_differentiableOn (hs : UniqueDiffOn 𝕜 s) : +theorem contDiffOn_iff_continuousOn_differentiableOn {n : ℕ∞} (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 := - ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin hm hs, fun _m hm => - h.differentiableOn_iteratedFDerivWithin hm hs⟩, + ⟨fun h => ⟨fun _m hm => h.continuousOn_iteratedFDerivWithin (mod_cast hm) hs, fun _m hm => + h.differentiableOn_iteratedFDerivWithin (mod_cast hm) hs⟩, fun h => contDiffOn_of_continuousOn_differentiableOn h.1 h.2⟩ theorem contDiffOn_succ_of_fderivWithin {n : ℕ} (hf : DifferentiableOn 𝕜 f s) @@ -628,7 +639,7 @@ theorem contDiffOn_succ_iff_fderivWithin {n : ℕ} (hs : UniqueDiffOn 𝕜 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 => ?_⟩ + refine ⟨H.differentiableOn le_add_self, fun x hx => ?_⟩ rcases contDiffWithinAt_succ_iff_hasFDerivWithinAt.1 (H x hx) with ⟨u, hu, f', hff', hf'⟩ rcases mem_nhdsWithin.1 hu with ⟨o, o_open, xo, ho⟩ rw [inter_comm, insert_eq_of_mem hx] at ho @@ -666,14 +677,14 @@ theorem contDiffOn_top_iff_fderivWithin (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 f s ∧ ContDiffOn 𝕜 ∞ (fun y => fderivWithin 𝕜 f s y) s := by constructor · intro h - refine ⟨h.differentiableOn le_top, ?_⟩ + refine ⟨h.differentiableOn (mod_cast le_top), ?_⟩ refine contDiffOn_top.2 fun n => ((contDiffOn_succ_iff_fderivWithin hs).1 ?_).2 - exact h.of_le le_top + exact h.of_le (mod_cast le_top) · intro h refine contDiffOn_top.2 fun n => ?_ - have A : (n : ℕ∞) ≤ ∞ := le_top + have A : (n : ℕ∞) ≤ ∞ := mod_cast le_top apply ((contDiffOn_succ_iff_fderivWithin hs).2 ⟨h.1, h.2.of_le A⟩).of_le - exact WithTop.coe_le_coe.2 (Nat.le_succ n) + exact_mod_cast (Nat.le_succ n) /-- A function is `C^∞` on an open domain if and only if it is differentiable there, and its derivative (expressed with `fderiv`) is `C^∞`. -/ @@ -684,13 +695,17 @@ theorem contDiffOn_top_iff_fderiv_of_isOpen (hs : IsOpen s) : protected theorem ContDiffOn.fderivWithin (hf : ContDiffOn 𝕜 n f s) (hs : UniqueDiffOn 𝕜 s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderivWithin 𝕜 f s y) s := by - cases' m with m - · change ∞ + 1 ≤ n at hmn - have : n = ∞ := by simpa using hmn - rw [this] at hf - exact ((contDiffOn_top_iff_fderivWithin hs).1 hf).2 - · change (m.succ : ℕ∞) ≤ n at hmn - exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 + rcases le_or_lt ∞ n with hn | hn + · have : ContDiffOn 𝕜 ∞ (fun y ↦ fderivWithin 𝕜 f s y) s := + ((contDiffOn_top_iff_fderivWithin hs).1 (hf.of_le hn)).2 + intro x hx k hk + exact this x hx k (mod_cast le_top) + · match m with + | ω => simpa using hmn.trans_lt hn + | ∞ => simpa using hmn.trans_lt hn + | (m : ℕ) => + change (m.succ : ℕ∞) ≤ n at hmn + exact ((contDiffOn_succ_iff_fderivWithin hs).1 (hf.of_le hmn)).2 theorem ContDiffOn.fderiv_of_isOpen (hf : ContDiffOn 𝕜 n f s) (hs : IsOpen s) (hmn : m + 1 ≤ n) : ContDiffOn 𝕜 m (fun y => fderiv 𝕜 f y) s := @@ -704,6 +719,13 @@ theorem ContDiffOn.continuousOn_fderiv_of_isOpen (h : ContDiffOn 𝕜 n f s) (hs (hn : 1 ≤ n) : ContinuousOn (fun x => fderiv 𝕜 f x) s := ((contDiffOn_succ_iff_fderiv_of_isOpen hs).1 (h.of_le hn)).2.continuousOn +/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, +it is only there as a convenient shortcut. -/ +theorem contDiffOn_infty_iff_contDiffOn_omega : + ContDiffOn 𝕜 ∞ f s ↔ ContDiffOn 𝕜 ω f s := by + have A (m : ℕ) : m ≤ ∞ := mod_cast le_top + simp [ContDiffOn, ContDiffWithinAt, hasFTaylorSeriesUpTo_top_iff, A] + /-! ### Smooth functions at a point -/ variable (𝕜) @@ -711,7 +733,7 @@ variable (𝕜) /-- A function is continuously differentiable up to `n` at a point `x` if, for any integer `k ≤ n`, there is a neighborhood of `x` where `f` admits derivatives up to order `n`, which are continuous. -/ -def ContDiffAt (n : ℕ∞) (f : E → F) (x : E) : Prop := +def ContDiffAt (n : WithTop ℕ∞) (f : E → F) (x : E) : Prop := ContDiffWithinAt 𝕜 n f univ x variable {𝕜} @@ -795,7 +817,7 @@ variable (𝕜) order `n`, which are continuous. Contrary to the case of definitions in domains (where derivatives might not be unique) we do not need to localize the definition in space or time. -/ -def ContDiff (n : ℕ∞) (f : E → F) : Prop := +def ContDiff (n : WithTop ℕ∞) (f : E → F) : Prop := ∃ p : E → FormalMultilinearSeries 𝕜 E F, HasFTaylorSeriesUpTo n f p variable {𝕜} @@ -824,10 +846,16 @@ theorem ContDiff.contDiffAt (h : ContDiff 𝕜 n f) : ContDiffAt 𝕜 n f x := theorem ContDiff.contDiffWithinAt (h : ContDiff 𝕜 n f) : ContDiffWithinAt 𝕜 n f s x := h.contDiffAt.contDiffWithinAt +/-- The following lemma will be removed when the definition of `C^ω` will be corrected. For now, +it is only there as a convenient shortcut. -/ +theorem contDiff_infty_iff_contDiff_omega : + ContDiff 𝕜 ∞ f ↔ ContDiff 𝕜 ω f := by + simp [ContDiff, hasFTaylorSeriesUpTo_top_iff] + theorem contDiff_top : ContDiff 𝕜 ∞ f ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by simp [contDiffOn_univ.symm, contDiffOn_top] -theorem contDiff_all_iff_nat : (∀ n, ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by +theorem contDiff_all_iff_nat : (∀ (n : ℕ∞), ContDiff 𝕜 n f) ↔ ∀ n : ℕ, ContDiff 𝕜 n f := by simp only [← contDiffOn_univ, contDiffOn_all_iff_nat] theorem ContDiff.contDiffOn (h : ContDiff 𝕜 n f) : ContDiffOn 𝕜 n f s := @@ -844,8 +872,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 - rw [show (1 : ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] - simp_rw [show ((0 : ℕ) : ℕ∞) = 0 from rfl, contDiffAt_zero, + rw [show (1 : WithTop ℕ∞) = (0 : ℕ) + 1 from rfl, contDiffAt_succ_iff_hasFDerivAt] + simp_rw [show ((0 : ℕ) : WithTop ℕ∞) = 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 +892,8 @@ theorem ContDiff.continuous (h : ContDiff 𝕜 n f) : Continuous f := theorem ContDiff.differentiable (h : ContDiff 𝕜 n f) (hn : 1 ≤ n) : Differentiable 𝕜 f := differentiableOn_univ.1 <| (contDiffOn_univ.2 h).differentiableOn hn -theorem contDiff_iff_forall_nat_le : ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → ContDiff 𝕜 m f := by +theorem contDiff_iff_forall_nat_le {n : ℕ∞} : + ContDiff 𝕜 n f ↔ ∀ m : ℕ, ↑m ≤ n → ContDiff 𝕜 m f := by simp_rw [← contDiffOn_univ]; exact contDiffOn_iff_forall_nat_le /-- A function is `C^(n+1)` iff it has a `C^n` derivative. -/ @@ -887,7 +916,7 @@ theorem contDiff_iff_ftaylorSeries : exact fun h => ContDiffOn.ftaylorSeriesWithin h uniqueDiffOn_univ · intro h; exact ⟨ftaylorSeries 𝕜 f, h⟩ -theorem contDiff_iff_continuous_differentiable : +theorem contDiff_iff_continuous_differentiable {n : ℕ∞} : ContDiff 𝕜 n f ↔ (∀ m : ℕ, m ≤ n → Continuous fun x => iteratedFDeriv 𝕜 m f x) ∧ ∀ m : ℕ, m < n → Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := by @@ -897,14 +926,15 @@ theorem contDiff_iff_continuous_differentiable : /-- 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) : Continuous fun x => iteratedFDeriv 𝕜 m f x := - (contDiff_iff_continuous_differentiable.mp hf).1 m hm + (contDiff_iff_continuous_differentiable.mp (hf.of_le hm)).1 m le_rfl /-- 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) : Differentiable 𝕜 fun x => iteratedFDeriv 𝕜 m f x := - (contDiff_iff_continuous_differentiable.mp hf).2 m hm + (contDiff_iff_continuous_differentiable.mp + (hf.of_le (ENat.add_one_natCast_le_withTop_of_lt hm))).2 m (mod_cast lt_add_one m) -theorem contDiff_of_differentiable_iteratedFDeriv +theorem contDiff_of_differentiable_iteratedFDeriv {n : ℕ∞} (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)⟩ diff --git a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean index 555d81bb107e4..c399175ddd00d 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/FiniteDimension.lean @@ -20,6 +20,7 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {D : Type uD} [NormedAddCommGroup D] [NormedSpace 𝕜 D] {E : Type uE} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {F : Type uF} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {n : WithTop ℕ∞} {f : D → E} {s : Set D} /-! ### Finite dimensional results -/ @@ -30,7 +31,7 @@ open Function Module variable [CompleteSpace 𝕜] /-- A family of continuous linear maps is `C^n` on `s` if all its applications are. -/ -theorem contDiffOn_clm_apply {n : ℕ∞} {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : +theorem contDiffOn_clm_apply {f : D → E →L[𝕜] F} {s : Set D} [FiniteDimensional 𝕜 E] : ContDiffOn 𝕜 n f s ↔ ∀ y, ContDiffOn 𝕜 n (fun x => f x y) s := by refine ⟨fun h y => h.clm_apply contDiffOn_const, fun h => ?_⟩ let d := finrank 𝕜 E @@ -40,7 +41,7 @@ theorem contDiffOn_clm_apply {n : ℕ∞} {f : D → E →L[𝕜] F} {s : Set D} rw [← id_comp f, ← e₂.symm_comp_self] exact e₂.symm.contDiff.comp_contDiffOn (contDiffOn_pi.mpr fun i => h _) -theorem contDiff_clm_apply_iff {n : ℕ∞} {f : D → E →L[𝕜] F} [FiniteDimensional 𝕜 E] : +theorem contDiff_clm_apply_iff {f : D → E →L[𝕜] F} [FiniteDimensional 𝕜 E] : ContDiff 𝕜 n f ↔ ∀ y, ContDiff 𝕜 n fun x => f x y := by simp_rw [← contDiffOn_univ, contDiffOn_clm_apply] @@ -48,7 +49,7 @@ theorem contDiff_clm_apply_iff {n : ℕ∞} {f : D → E →L[𝕜] F} [FiniteDi When you do induction on `n`, this gives a useful characterization of a function being `C^(n+1)`, assuming you have already computed the derivative. The advantage of this version over `contDiff_succ_iff_fderiv` is that both occurrences of `ContDiff` are for functions with the same -domain and codomain (`E` and `F`). This is not the case for `contDiff_succ_iff_fderiv`, which +domain and codomain (`D` and `E`). This is not the case for `contDiff_succ_iff_fderiv`, which often requires an inconvenient need to generalize `F`, which results in universe issues (see the discussion in the section of `ContDiff.comp`). diff --git a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean index 5f8fdba324bcb..e87af872c43a2 100644 --- a/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean +++ b/Mathlib/Analysis/Calculus/ContDiff/RCLike.lean @@ -24,8 +24,8 @@ section Real its extension fields such as `ℂ`). -/ -variable {n : ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕂 E'] - {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕂 F'] +variable {n : WithTop ℕ∞} {𝕂 : Type*} [RCLike 𝕂] {E' : Type*} [NormedAddCommGroup E'] + [NormedSpace 𝕂 E'] {F' : Type*} [NormedAddCommGroup F'] [NormedSpace 𝕂 F'] /-- 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`. -/ @@ -41,7 +41,7 @@ us as `f'`, then `f'` is also a strict derivative. -/ theorem ContDiffAt.hasStrictFDerivAt' {f : E' → F'} {f' : E' →L[𝕂] F'} {x : E'} (hf : ContDiffAt 𝕂 n f x) (hf' : HasFDerivAt f f' x) (hn : 1 ≤ n) : HasStrictFDerivAt f f' x := by - rcases hf 1 hn with ⟨u, H, p, hp⟩ + rcases hf.of_le hn 1 le_rfl with ⟨u, H, p, hp⟩ simp only [nhdsWithin_univ, mem_univ, insert_eq_of_mem] at H have := hp.hasStrictFDerivAt le_rfl H rwa [hf'.unique this.hasFDerivAt] @@ -135,7 +135,7 @@ lemma ContDiff.locallyLipschitz {f : E' → F'} (hf : ContDiff 𝕂 1 f) : Local use K, t /-- A `C^1` function with compact support is Lipschitz. -/ -theorem ContDiff.lipschitzWith_of_hasCompactSupport {f : E' → F'} {n : ℕ∞} +theorem ContDiff.lipschitzWith_of_hasCompactSupport {f : E' → F'} (hf : HasCompactSupport f) (h'f : ContDiff 𝕂 n f) (hn : 1 ≤ n) : ∃ C, LipschitzWith C f := by obtain ⟨C, hC⟩ := (hf.fderiv 𝕂).exists_bound_of_continuous (h'f.continuous_fderiv hn) diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 055c8cb25720e..9141ffab19f41 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -63,7 +63,7 @@ differentiability at points in a neighborhood of `s`. Therefore, the theorem tha open Filter Asymptotics Set -open scoped ENNReal Topology +open scoped ENNReal Topology ContDiff universe u v @@ -261,7 +261,7 @@ by the sequence of its derivatives. Note that, if the function were just analyti one would have to use instead the sequence of derivatives inside the set, as in `AnalyticOn.hasFTaylorSeriesUpToOn`. -/ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] - (n : ℕ∞) (h : AnalyticOnNhd 𝕜 f s) : + (n : WithTop ℕ∞) (h : AnalyticOnNhd 𝕜 f s) : HasFTaylorSeriesUpToOn n f (ftaylorSeries 𝕜 f) s := by refine ⟨fun x _hx ↦ rfl, fun m _hm x hx ↦ ?_, fun m _hm x hx ↦ ?_⟩ · apply HasFDerivAt.hasFDerivWithinAt @@ -270,25 +270,27 @@ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] exact (h.iteratedFDeriv m x hx).differentiableAt /-- An analytic function is infinitely differentiable. -/ -protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) {n : ℕ∞} : - ContDiffOn 𝕜 n f s := +protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) + {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top + rw [← contDiffOn_infty_iff_contDiffOn_omega] let t := { x | AnalyticAt 𝕜 f x } - suffices ContDiffOn 𝕜 n f t from this.mono h + suffices ContDiffOn 𝕜 ∞ f t from this.mono h have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx have t_open : IsOpen t := isOpen_analyticAt 𝕜 f - contDiffOn_of_continuousOn_differentiableOn + exact contDiffOn_of_continuousOn_differentiableOn (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) (fun m _ ↦ (H.iteratedFDeriv m).differentiableOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) /-- An analytic function on the whole space is infinitely differentiable there. -/ -theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : ℕ∞} : +theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : WithTop ℕ∞} : ContDiff 𝕜 n f := by rw [← contDiffOn_univ] exact h.contDiffOn -theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : ℕ∞} : +theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f x := by obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd exact hf.contDiffOn.contDiffAt hs @@ -306,7 +308,7 @@ protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn lemma AnalyticWithinAt.exists_hasFTaylorSeriesUpToOn [CompleteSpace F] - (n : ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : + (n : WithTop ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : ∃ u ∈ 𝓝[insert x s] x, ∃ (p : E → FormalMultilinearSeries 𝕜 E F), HasFTaylorSeriesUpToOn n f p u ∧ ∀ i, AnalyticOn 𝕜 (fun x ↦ p x i) u := by rcases h.exists_analyticAt with ⟨g, -, fg, hg⟩ @@ -545,19 +547,21 @@ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : simp /-- A polynomial function is infinitely differentiable. -/ -theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : ℕ∞} : - ContDiffOn 𝕜 n f s := +theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : + ContDiffOn 𝕜 n f s := by + suffices ContDiffOn 𝕜 ω f s from this.of_le le_top let t := { x | CPolynomialAt 𝕜 f x } - suffices ContDiffOn 𝕜 n f t from this.mono h + suffices ContDiffOn 𝕜 ω f t from this.mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f - contDiffOn_of_continuousOn_differentiableOn + exact contDiffOn_of_continuousOn_differentiableOn (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) -theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : ℕ∞} : +theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : ContDiffAt 𝕜 n f x := let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn hf.contDiffOn.contDiffAt hs @@ -595,7 +599,7 @@ theorem changeOriginSeries_support {k l : ℕ} (h : k + l ≠ Fintype.card ι) : simp_rw [FormalMultilinearSeries.changeOriginSeriesTerm, toFormalMultilinearSeries, dif_neg h.symm, LinearIsometryEquiv.map_zero] -variable {n : ℕ∞} (x : ∀ i, E i) +variable {n : WithTop ℕ∞} (x : ∀ i, E i) open Finset in theorem changeOrigin_toFormalMultilinearSeries [DecidableEq ι] : diff --git a/Mathlib/Analysis/Calculus/FDeriv/Norm.lean b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean index 4587f2bbf6dfc..fabdde33a4bfd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Norm.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Norm.lean @@ -38,7 +38,7 @@ differentiability, norm open ContinuousLinearMap Filter NNReal Real Set variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] -variable {n : ℕ∞} {f : E →L[ℝ] ℝ} {x : E} {t : ℝ} +variable {n : WithTop ℕ∞} {f : E →L[ℝ] ℝ} {x : E} {t : ℝ} variable (E) in theorem not_differentiableAt_norm_zero [Nontrivial E] : @@ -72,15 +72,15 @@ theorem contDiffAt_norm_smul_iff (ht : t ≠ 0) : theorem ContDiffAt.contDiffAt_norm_of_smul (h : ContDiffAt ℝ n (‖·‖) (t • x)) : ContDiffAt ℝ n (‖·‖) x := by - obtain rfl | hn : n = 0 ∨ 1 ≤ n := by - rw [← ENat.lt_one_iff_eq_zero] - exact lt_or_le .. - · rw [contDiffAt_zero] + rcases eq_bot_or_bot_lt n with rfl | hn + · apply contDiffAt_zero.2 exact ⟨univ, univ_mem, continuous_norm.continuousOn⟩ + replace hn : 1 ≤ n := ENat.add_one_natCast_le_withTop_of_lt hn obtain rfl | ht := eq_or_ne t 0 · by_cases hE : Nontrivial E · rw [zero_smul] at h - exact (mt (ContDiffAt.differentiableAt · hn)) (not_differentiableAt_norm_zero E) h |>.elim + exact (mt (ContDiffAt.differentiableAt · (mod_cast hn))) + (not_differentiableAt_norm_zero E) h |>.elim · rw [not_nontrivial_iff_subsingleton] at hE rw [eq_const_of_subsingleton (‖·‖) 0] exact contDiffAt_const diff --git a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean index a893dcc67f3d8..b4f9b4db76e68 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Symmetric.lean @@ -444,7 +444,7 @@ theorem second_derivative_symmetric {f' : E → E →L[𝕜] F} {f'' : E →L[ second_derivative_symmetric_of_eventually (Filter.Eventually.of_forall hf) hx v w /-- If a function is `C^2` at a point, then its second derivative there is symmetric. -/ -theorem ContDiffAt.isSymmSndFDerivAt {n : ℕ∞} (hf : ContDiffAt 𝕜 n f x) (hn : 2 ≤ n) : +theorem ContDiffAt.isSymmSndFDerivAt {n : WithTop ℕ∞} (hf : ContDiffAt 𝕜 n f x) (hn : 2 ≤ n) : IsSymmSndFDerivAt 𝕜 f x := by intro v w apply second_derivative_symmetric_of_eventually (f := f) (f' := fderiv 𝕜 f) (x := x) @@ -462,7 +462,7 @@ theorem ContDiffAt.isSymmSndFDerivAt {n : ℕ∞} (hf : ContDiffAt 𝕜 n f x) ( /-- If a function is `C^2` within a set at a point, and accumulated by points in the interior of the set, then its second derivative there is symmetric. -/ -theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : ℕ∞} (hf : ContDiffWithinAt 𝕜 n f s x) +theorem ContDiffWithinAt.isSymmSndFDerivWithinAt {n : WithTop ℕ∞} (hf : ContDiffWithinAt 𝕜 n f s x) (hn : 2 ≤ n) (hs : UniqueDiffOn 𝕜 s) (hx : x ∈ closure (interior s)) (h'x : x ∈ s) : IsSymmSndFDerivWithinAt 𝕜 f s x := by /- We argue that, at interior points, the second derivative is symmetric, and moreover by diff --git a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean index 55e18e301150a..49b5c3a3c929c 100644 --- a/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean +++ b/Mathlib/Analysis/Calculus/InverseFunctionTheorem/ContDiff.lean @@ -20,46 +20,46 @@ namespace ContDiffAt variable {𝕂 : Type*} [RCLike 𝕂] variable {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕂 E] variable {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕂 F] -variable [CompleteSpace E] (f : E → F) {f' : E ≃L[𝕂] F} {a : E} +variable [CompleteSpace E] (f : E → F) {f' : E ≃L[𝕂] F} {a : E} {n : WithTop ℕ∞} /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a `PartialHomeomorph` with `to_fun = f` and `a ∈ source`. -/ -def toPartialHomeomorph {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) +def toPartialHomeomorph (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : PartialHomeomorph E F := (hf.hasStrictFDerivAt' hf' hn).toPartialHomeomorph f variable {f} @[simp] -theorem toPartialHomeomorph_coe {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem toPartialHomeomorph_coe (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : (hf.toPartialHomeomorph f hf' hn : E → F) = f := rfl -theorem mem_toPartialHomeomorph_source {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem mem_toPartialHomeomorph_source (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : a ∈ (hf.toPartialHomeomorph f hf' hn).source := (hf.hasStrictFDerivAt' hf' hn).mem_toPartialHomeomorph_source -theorem image_mem_toPartialHomeomorph_target {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem image_mem_toPartialHomeomorph_target (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : f a ∈ (hf.toPartialHomeomorph f hf' hn).target := (hf.hasStrictFDerivAt' hf' hn).image_mem_toPartialHomeomorph_target /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, returns a function that is locally inverse to `f`. -/ -def localInverse {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) +def localInverse (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : F → E := (hf.hasStrictFDerivAt' hf' hn).localInverse f f' a -theorem localInverse_apply_image {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem localInverse_apply_image (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : hf.localInverse hf' hn (f a) = a := (hf.hasStrictFDerivAt' hf' hn).localInverse_apply_image /-- Given a `ContDiff` function over `𝕂` (which is `ℝ` or `ℂ`) with an invertible derivative at `a`, the inverse function (produced by `ContDiff.toPartialHomeomorph`) is also `ContDiff`. -/ -theorem to_localInverse {n : ℕ∞} (hf : ContDiffAt 𝕂 n f a) +theorem to_localInverse (hf : ContDiffAt 𝕂 n f a) (hf' : HasFDerivAt f (f' : E →L[𝕂] F) a) (hn : 1 ≤ n) : ContDiffAt 𝕂 n (hf.localInverse hf' hn) (f a) := by have := hf.localInverse_apply_image hf' hn diff --git a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean index 9fcaa3d13a0cc..858642c83caa8 100644 --- a/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean +++ b/Mathlib/Analysis/Calculus/IteratedDeriv/Defs.lean @@ -131,12 +131,13 @@ theorem contDiffOn_of_differentiableOn_deriv {n : ℕ∞} /-- On a set with unique derivatives, a `C^n` function has derivatives up to `n` which are continuous. -/ -theorem ContDiffOn.continuousOn_iteratedDerivWithin {n : ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) +theorem ContDiffOn.continuousOn_iteratedDerivWithin + {n : WithTop ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) (hmn : (m : ℕ∞) ≤ n) (hs : UniqueDiffOn 𝕜 s) : ContinuousOn (iteratedDerivWithin m f s) s := by simpa only [iteratedDerivWithin_eq_equiv_comp, LinearIsometryEquiv.comp_continuousOn_iff] using h.continuousOn_iteratedFDerivWithin hmn hs -theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : ℕ∞} {m : ℕ} +theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} (h : ContDiffWithinAt 𝕜 n f s x) (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 (insert x s)) : DifferentiableWithinAt 𝕜 (iteratedDerivWithin m f s) s x := by simpa only [iteratedDerivWithin_eq_equiv_comp, @@ -145,8 +146,8 @@ theorem ContDiffWithinAt.differentiableWithinAt_iteratedDerivWithin {n : ℕ∞} /-- On a set with unique derivatives, a `C^n` function has derivatives less than `n` which are differentiable. -/ -theorem ContDiffOn.differentiableOn_iteratedDerivWithin {n : ℕ∞} {m : ℕ} (h : ContDiffOn 𝕜 n f s) - (hmn : (m : ℕ∞) < n) (hs : UniqueDiffOn 𝕜 s) : +theorem ContDiffOn.differentiableOn_iteratedDerivWithin {n : WithTop ℕ∞} {m : ℕ} + (h : ContDiffOn 𝕜 n f s) (hmn : m < n) (hs : UniqueDiffOn 𝕜 s) : DifferentiableOn 𝕜 (iteratedDerivWithin m f s) s := fun x hx => (h x hx).differentiableWithinAt_iteratedDerivWithin hmn <| by rwa [insert_eq_of_mem hx] @@ -238,13 +239,14 @@ theorem contDiff_of_differentiable_iteratedDeriv {n : ℕ∞} (h : ∀ m : ℕ, (m : ℕ∞) ≤ n → Differentiable 𝕜 (iteratedDeriv m f)) : ContDiff 𝕜 n f := contDiff_iff_iteratedDeriv.2 ⟨fun m hm => (h m hm).continuous, fun m hm => h m (le_of_lt hm)⟩ -theorem ContDiff.continuous_iteratedDeriv {n : ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) - (hmn : (m : ℕ∞) ≤ n) : Continuous (iteratedDeriv m f) := - (contDiff_iff_iteratedDeriv.1 h).1 m hmn +theorem ContDiff.continuous_iteratedDeriv {n : WithTop ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) + (hmn : m ≤ n) : Continuous (iteratedDeriv m f) := + (contDiff_iff_iteratedDeriv.1 (h.of_le hmn)).1 m le_rfl -theorem ContDiff.differentiable_iteratedDeriv {n : ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) - (hmn : (m : ℕ∞) < n) : Differentiable 𝕜 (iteratedDeriv m f) := - (contDiff_iff_iteratedDeriv.1 h).2 m hmn +theorem ContDiff.differentiable_iteratedDeriv {n : WithTop ℕ∞} (m : ℕ) (h : ContDiff 𝕜 n f) + (hmn : m < n) : Differentiable 𝕜 (iteratedDeriv m f) := + (contDiff_iff_iteratedDeriv.1 (h.of_le (ENat.add_one_natCast_le_withTop_of_lt hmn))).2 m + (mod_cast (lt_add_one m)) /-- The `n+1`-th iterated derivative can be obtained by differentiating the `n`-th iterated derivative. -/ diff --git a/Mathlib/Analysis/Calculus/Rademacher.lean b/Mathlib/Analysis/Calculus/Rademacher.lean index 28ad0559741a1..1c836a3a36f92 100644 --- a/Mathlib/Analysis/Calculus/Rademacher.lean +++ b/Mathlib/Analysis/Calculus/Rademacher.lean @@ -221,9 +221,9 @@ theorem ae_lineDeriv_sum_eq suffices S2 : ∫ x, (∑ i ∈ s, a i * fderiv ℝ g x (v i)) * f x ∂μ = ∑ i ∈ s, a i * ∫ x, fderiv ℝ g x (v i) * f x ∂μ by obtain ⟨D, g_lip⟩ : ∃ D, LipschitzWith D g := - ContDiff.lipschitzWith_of_hasCompactSupport g_comp g_smooth le_top + ContDiff.lipschitzWith_of_hasCompactSupport g_comp g_smooth (mod_cast le_top) simp_rw [integral_lineDeriv_mul_eq hf g_lip g_comp] - simp_rw [(g_smooth.differentiable le_top).differentiableAt.lineDeriv_eq_fderiv] + simp_rw [(g_smooth.differentiable (mod_cast le_top)).differentiableAt.lineDeriv_eq_fderiv] simp only [map_neg, _root_.map_sum, _root_.map_smul, smul_eq_mul, neg_mul] simp only [integral_neg, mul_neg, Finset.sum_neg_distrib, neg_inj] exact S2 @@ -233,7 +233,8 @@ theorem ae_lineDeriv_sum_eq let L : (E →L[ℝ] ℝ) → ℝ := fun f ↦ f (v i) change Integrable (fun x ↦ a i * ((L ∘ (fderiv ℝ g)) x * f x)) μ refine (Continuous.integrable_of_hasCompactSupport ?_ ?_).const_mul _ - · exact ((g_smooth.continuous_fderiv le_top).clm_apply continuous_const).mul hf.continuous + · exact ((g_smooth.continuous_fderiv (mod_cast le_top)).clm_apply continuous_const).mul + hf.continuous · exact ((g_comp.fderiv ℝ).comp_left rfl).mul_right /-! diff --git a/Mathlib/Analysis/Calculus/SmoothSeries.lean b/Mathlib/Analysis/Calculus/SmoothSeries.lean index 58d0f39d86c54..073a4b2192939 100644 --- a/Mathlib/Analysis/Calculus/SmoothSeries.lean +++ b/Mathlib/Analysis/Calculus/SmoothSeries.lean @@ -198,7 +198,8 @@ theorem iteratedFDeriv_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) have A : Summable fun n => iteratedFDeriv 𝕜 k (f n) 0 := .of_norm_bounded (v k) (hv k h'k.le) fun n => h'f k n 0 h'k.le simp_rw [iteratedFDeriv_succ_eq_comp_left, IH h'k.le] - rw [fderiv_tsum (hv _ hk) (fun n => (hf n).differentiable_iteratedFDeriv h'k) _ A] + rw [fderiv_tsum (hv _ hk) (fun n => (hf n).differentiable_iteratedFDeriv + (mod_cast h'k)) _ A] · ext1 x exact (continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (k + 1) => E) F).symm.toContinuousLinearEquiv.map_tsum @@ -219,7 +220,7 @@ theorem iteratedFDeriv_tsum_apply (hf : ∀ i, ContDiff 𝕜 N (f i)) class `C^N`, and moreover there is a uniform summable upper bound on the `k`-th derivative for each `k ≤ N`. Then the series is also `C^N`. -/ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k : ℕ∞) ≤ N → Summable (v k)) - (h'f : ∀ (k : ℕ) (i : α) (x : E), (k : ℕ∞) ≤ N → ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : + (h'f : ∀ (k : ℕ) (i : α) (x : E), k ≤ N → ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : ContDiff 𝕜 N fun x => ∑' i, f i x := by rw [contDiff_iff_continuous_differentiable] constructor @@ -227,16 +228,16 @@ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k rw [iteratedFDeriv_tsum hf hv h'f hm] refine continuous_tsum ?_ (hv m hm) ?_ · intro i - exact ContDiff.continuous_iteratedFDeriv hm (hf i) + exact ContDiff.continuous_iteratedFDeriv (mod_cast hm) (hf i) · intro n x exact h'f _ _ _ hm · intro m hm have h'm : ((m + 1 : ℕ) : ℕ∞) ≤ N := by simpa only [ENat.coe_add, ENat.coe_one] using Order.add_one_le_of_lt hm rw [iteratedFDeriv_tsum hf hv h'f hm.le] - have A : - ∀ n x, HasFDerivAt (iteratedFDeriv 𝕜 m (f n)) (fderiv 𝕜 (iteratedFDeriv 𝕜 m (f n)) x) x := - fun n x => (ContDiff.differentiable_iteratedFDeriv hm (hf n)).differentiableAt.hasFDerivAt + have A n x : HasFDerivAt (iteratedFDeriv 𝕜 m (f n)) (fderiv 𝕜 (iteratedFDeriv 𝕜 m (f n)) x) x := + (ContDiff.differentiable_iteratedFDeriv (mod_cast hm) + (hf n)).differentiableAt.hasFDerivAt refine differentiable_tsum (hv _ h'm) A fun n x => ?_ rw [fderiv_iteratedFDeriv, comp_apply, LinearIsometryEquiv.norm_map] exact h'f _ _ _ h'm @@ -245,11 +246,9 @@ theorem contDiff_tsum (hf : ∀ i, ContDiff 𝕜 N (f i)) (hv : ∀ k : ℕ, (k class `C^N`, and moreover there is a uniform summable upper bound on the `k`-th derivative for each `k ≤ N` (except maybe for finitely many `i`s). Then the series is also `C^N`. -/ theorem contDiff_tsum_of_eventually (hf : ∀ i, ContDiff 𝕜 N (f i)) - (hv : ∀ k : ℕ, (k : ℕ∞) ≤ N → Summable (v k)) - (h'f : - ∀ k : ℕ, - (k : ℕ∞) ≤ N → - ∀ᶠ i in (Filter.cofinite : Filter α), ∀ x : E, ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : + (hv : ∀ k : ℕ, k ≤ N → Summable (v k)) + (h'f : ∀ k : ℕ, k ≤ N → + ∀ᶠ i in (Filter.cofinite : Filter α), ∀ x : E, ‖iteratedFDeriv 𝕜 k (f i) x‖ ≤ v k i) : ContDiff 𝕜 N fun x => ∑' i, f i x := by classical refine contDiff_iff_forall_nat_le.2 fun m hm => ?_ @@ -274,10 +273,10 @@ theorem contDiff_tsum_of_eventually (hf : ∀ i, ContDiff 𝕜 N (f i)) filter_upwards [h'f 0 (zero_le _)] with i hi simpa only [norm_iteratedFDeriv_zero] using hi x rw [this] - apply (ContDiff.sum fun i _ => (hf i).of_le hm).add + apply (ContDiff.sum fun i _ => (hf i).of_le (mod_cast hm)).add have h'u : ∀ k : ℕ, (k : ℕ∞) ≤ m → Summable (v k ∘ ((↑) : { i // i ∉ T } → α)) := fun k hk => (hv k (hk.trans hm)).subtype _ - refine contDiff_tsum (fun i => (hf i).of_le hm) h'u ?_ + refine contDiff_tsum (fun i => (hf i).of_le (mod_cast hm)) h'u ?_ rintro k ⟨i, hi⟩ x hk simp only [t, T, Finite.mem_toFinset, mem_setOf_eq, Finset.mem_range, not_forall, not_le, exists_prop, not_exists, not_and, not_lt] at hi diff --git a/Mathlib/Analysis/Calculus/Taylor.lean b/Mathlib/Analysis/Calculus/Taylor.lean index 71293eb9d3058..8a56fcf0d80a4 100644 --- a/Mathlib/Analysis/Calculus/Taylor.lean +++ b/Mathlib/Analysis/Calculus/Taylor.lean @@ -119,7 +119,8 @@ theorem continuousOn_taylorWithinEval {f : ℝ → E} {x : ℝ} {n : ℕ} {s : S simp_rw [taylor_within_apply] refine continuousOn_finset_sum (Finset.range (n + 1)) fun i hi => ?_ refine (continuousOn_const.mul ((continuousOn_const.sub continuousOn_id).pow _)).smul ?_ - rw [contDiffOn_iff_continuousOn_differentiableOn_deriv hs] at hf + rw [show (n : WithTop ℕ∞) = (n : ℕ∞) by rfl, + contDiffOn_iff_continuousOn_differentiableOn_deriv hs] at hf cases' hf with hf_left specialize hf_left i simp only [Finset.mem_range] at hi @@ -179,7 +180,7 @@ theorem hasDerivWithinAt_taylorWithinEval {f : ℝ → E} {x y : ℝ} {n : ℕ} simp only [add_zero, Nat.factorial_succ, Nat.cast_mul, Nat.cast_add, Nat.cast_one] have coe_lt_succ : (k : WithTop ℕ) < k.succ := Nat.cast_lt.2 k.lt_succ_self have hdiff : DifferentiableOn ℝ (iteratedDerivWithin k f s) s' := - (hf.differentiableOn_iteratedDerivWithin coe_lt_succ hs_unique).mono h + (hf.differentiableOn_iteratedDerivWithin (mod_cast coe_lt_succ) hs_unique).mono h specialize hk hf.of_succ ((hdiff y hy).mono_of_mem_nhdsWithin hs') convert hk.add (hasDerivWithinAt_taylor_coeff_within hs'_unique (nhdsWithin_mono _ h self_mem_nhdsWithin) hf') using 1 @@ -305,7 +306,7 @@ theorem taylor_mean_remainder_bound {f : ℝ → E} {a b C x : ℝ} {n : ℕ} (h simp [hx] -- The nth iterated derivative is differentiable have hf' : DifferentiableOn ℝ (iteratedDerivWithin n f (Icc a b)) (Icc a b) := - hf.differentiableOn_iteratedDerivWithin (WithTop.coe_lt_coe.mpr n.lt_succ_self) + hf.differentiableOn_iteratedDerivWithin (mod_cast n.lt_succ_self) (uniqueDiffOn_Icc h) -- We can uniformly bound the derivative of the Taylor polynomial have h' : ∀ y ∈ Ico a x, diff --git a/Mathlib/Analysis/Calculus/VectorField.lean b/Mathlib/Analysis/Calculus/VectorField.lean index d4e54bc4485f9..ac4eb32fdc707 100644 --- a/Mathlib/Analysis/Calculus/VectorField.lean +++ b/Mathlib/Analysis/Calculus/VectorField.lean @@ -144,7 +144,7 @@ lemma lieBracket_swap : lieBracket 𝕜 V W x = - lieBracket 𝕜 W V x := by ext x; simp [lieBracket] lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField - {m n : ℕ∞} (hV : ContDiffWithinAt 𝕜 n V s x) + {m n : WithTop ℕ∞} (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 @@ -153,19 +153,19 @@ lemma _root_.ContDiffWithinAt.lieBracketWithin_vectorField · 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) +lemma _root_.ContDiffAt.lieBracket_vectorField {m n : WithTop ℕ∞} (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) +lemma _root_.ContDiffOn.lieBracketWithin_vectorField {m n : WithTop ℕ∞} (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) +lemma _root_.ContDiff.lieBracket_vectorField {m n : WithTop ℕ∞} (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) diff --git a/Mathlib/Analysis/Complex/RealDeriv.lean b/Mathlib/Analysis/Complex/RealDeriv.lean index 3ce85a20eb6f5..1609fa5e357a4 100644 --- a/Mathlib/Analysis/Complex/RealDeriv.lean +++ b/Mathlib/Analysis/Complex/RealDeriv.lean @@ -77,14 +77,14 @@ theorem HasDerivAt.real_of_complex (h : HasDerivAt e e' z) : rw [ContinuousLinearMap.comp_apply, ContinuousLinearMap.comp_apply] simp -theorem ContDiffAt.real_of_complex {n : ℕ∞} (h : ContDiffAt ℂ n e z) : +theorem ContDiffAt.real_of_complex {n : WithTop ℕ∞} (h : ContDiffAt ℂ n e z) : ContDiffAt ℝ n (fun x : ℝ => (e x).re) z := by have A : ContDiffAt ℝ n ((↑) : ℝ → ℂ) z := ofRealCLM.contDiff.contDiffAt have B : ContDiffAt ℝ n e z := h.restrict_scalars ℝ have C : ContDiffAt ℝ n re (e z) := reCLM.contDiff.contDiffAt exact C.comp z (B.comp z A) -theorem ContDiff.real_of_complex {n : ℕ∞} (h : ContDiff ℂ n e) : +theorem ContDiff.real_of_complex {n : WithTop ℕ∞} (h : ContDiff ℂ n e) : ContDiff ℝ n fun x : ℝ => (e x).re := contDiff_iff_contDiffAt.2 fun _ => h.contDiffAt.real_of_complex diff --git a/Mathlib/Analysis/Convolution.lean b/Mathlib/Analysis/Convolution.lean index e3a80a3cd256e..f2e6fd1a550aa 100644 --- a/Mathlib/Analysis/Convolution.lean +++ b/Mathlib/Analysis/Convolution.lean @@ -1179,17 +1179,18 @@ theorem contDiffOn_convolution_right_with_param_aux {G : Type uP} {E' : Type uP} come from the same universe). -/ induction n using ENat.nat_induction generalizing g E' F with | h0 => - rw [contDiffOn_zero] at hg ⊢ + rw [WithTop.coe_zero, contDiffOn_zero] at hg ⊢ exact continuousOn_convolution_right_with_param L hk hgs hf hg | hsuc n ih => + simp only [Nat.succ_eq_add_one, Nat.cast_add, Nat.cast_one, WithTop.coe_add, + WithTop.coe_natCast, WithTop.coe_one] at hg ⊢ let f' : P → G → P × G →L[𝕜] F := fun p a => (f ⋆[L.precompR (P × G), μ] fun x : G => fderiv 𝕜 (uncurry g) (p, x)) a 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 [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, - contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ - constructor + rw [contDiffOn_succ_iff_fderiv_of_isOpen (hs.prod (@isOpen_univ G _))] at hg ⊢ + refine ⟨?_, ?_⟩ · rintro ⟨p, x⟩ ⟨hp, -⟩ exact (A (p, x) hp).differentiableAt.differentiableWithinAt · suffices H : ContDiffOn 𝕜 n (↿f') (s ×ˢ univ) by diff --git a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean index b6d2c67574d34..8aa816c388f41 100644 --- a/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean +++ b/Mathlib/Analysis/Distribution/AEEqOfIntegralContDiff.lean @@ -24,7 +24,7 @@ as `ae_eq_zero_of_integral_smooth_smul_eq_zero` and `ae_eq_of_integral_smooth_sm open MeasureTheory Filter Metric Function Set TopologicalSpace -open scoped Topology Manifold +open scoped Topology Manifold ContDiff variable {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] [FiniteDimensional ℝ E] {F : Type*} [NormedAddCommGroup F] [NormedSpace ℝ F] [CompleteSpace F] @@ -182,7 +182,7 @@ variable [MeasurableSpace E] [BorelSpace E] {f f' : E → F} {μ : Measure E} /-- If a locally integrable function `f` on a finite-dimensional real vector space has zero integral when multiplied by any smooth compactly supported function, then `f` vanishes almost everywhere. -/ theorem ae_eq_zero_of_integral_contDiff_smul_eq_zero (hf : LocallyIntegrable f μ) - (h : ∀ (g : E → ℝ), ContDiff ℝ ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : + (h : ∀ (g : E → ℝ), ContDiff ℝ ∞ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, f x = 0 := ae_eq_zero_of_integral_smooth_smul_eq_zero 𝓘(ℝ, E) hf (fun g g_diff g_supp ↦ h g g_diff.contDiff g_supp) @@ -192,7 +192,7 @@ integral when multiplied by any smooth compactly supported function, then they c everywhere. -/ theorem ae_eq_of_integral_contDiff_smul_eq (hf : LocallyIntegrable f μ) (hf' : LocallyIntegrable f' μ) (h : ∀ (g : E → ℝ), - ContDiff ℝ ⊤ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : + ContDiff ℝ ∞ g → HasCompactSupport g → ∫ x, g x • f x ∂μ = ∫ x, g x • f' x ∂μ) : ∀ᵐ x ∂μ, f x = f' x := ae_eq_of_integral_smooth_smul_eq 𝓘(ℝ, E) hf hf' (fun g g_diff g_supp ↦ h g g_diff.contDiff g_supp) @@ -202,7 +202,7 @@ theorem ae_eq_of_integral_contDiff_smul_eq in an open set `U`, then `f` vanishes almost everywhere in `U`. -/ theorem IsOpen.ae_eq_zero_of_integral_contDiff_smul_eq_zero {U : Set E} (hU : IsOpen U) (hf : LocallyIntegrableOn f U μ) - (h : ∀ (g : E → ℝ), ContDiff ℝ ⊤ g → HasCompactSupport g → tsupport g ⊆ U → + (h : ∀ (g : E → ℝ), ContDiff ℝ ∞ g → HasCompactSupport g → tsupport g ⊆ U → ∫ x, g x • f x ∂μ = 0) : ∀ᵐ x ∂μ, x ∈ U → f x = 0 := hU.ae_eq_zero_of_integral_smooth_smul_eq_zero 𝓘(ℝ, E) hf diff --git a/Mathlib/Analysis/Distribution/SchwartzSpace.lean b/Mathlib/Analysis/Distribution/SchwartzSpace.lean index 7aab3fe966d3b..b1f19e08a826b 100644 --- a/Mathlib/Analysis/Distribution/SchwartzSpace.lean +++ b/Mathlib/Analysis/Distribution/SchwartzSpace.lean @@ -61,7 +61,7 @@ Schwartz space, tempered distributions noncomputable section -open scoped Nat NNReal +open scoped Nat NNReal ContDiff variable {𝕜 𝕜' D E F G V : Type*} variable [NormedAddCommGroup E] [NormedSpace ℝ E] @@ -72,7 +72,7 @@ variable (E F) any power of `‖x‖`. -/ structure SchwartzMap where toFun : E → F - smooth' : ContDiff ℝ ⊤ toFun + smooth' : ContDiff ℝ ∞ toFun decay' : ∀ k n : ℕ, ∃ C : ℝ, ∀ x, ‖x‖ ^ k * ‖iteratedFDeriv ℝ n toFun x‖ ≤ C /-- A function is a Schwartz function if it is smooth and all derivatives decay faster than @@ -98,7 +98,7 @@ theorem decay (f : 𝓢(E, F)) (k n : ℕ) : /-- Every Schwartz function is smooth. -/ theorem smooth (f : 𝓢(E, F)) (n : ℕ∞) : ContDiff ℝ n f := - f.smooth'.of_le le_top + f.smooth'.of_le (mod_cast le_top) /-- Every Schwartz function is continuous. -/ @[continuity] @@ -521,7 +521,7 @@ section TemperateGrowth /-- A function is called of temperate growth if it is smooth and all iterated derivatives are polynomially bounded. -/ def _root_.Function.HasTemperateGrowth (f : E → F) : Prop := - ContDiff ℝ ⊤ f ∧ ∀ n : ℕ, ∃ (k : ℕ) (C : ℝ), ∀ x, ‖iteratedFDeriv ℝ n f x‖ ≤ C * (1 + ‖x‖) ^ k + ContDiff ℝ ∞ f ∧ ∀ n : ℕ, ∃ (k : ℕ) (C : ℝ), ∀ x, ‖iteratedFDeriv ℝ n f x‖ ≤ C * (1 + ‖x‖) ^ k theorem _root_.Function.HasTemperateGrowth.norm_iteratedFDeriv_le_uniform_aux {f : E → F} (hf_temperate : f.HasTemperateGrowth) (n : ℕ) : @@ -680,7 +680,7 @@ variable {σ : 𝕜 →+* 𝕜'} Note: This is a helper definition for `mkCLM`. -/ def mkLM (A : (D → E) → F → G) (hadd : ∀ (f g : 𝓢(D, E)) (x), A (f + g) x = A f x + A g x) (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) (x), A (a • f) x = σ a • A f x) - (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ⊤ (A f)) + (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ∞ (A f)) (hbound : ∀ n : ℕ × ℕ, ∃ (s : Finset (ℕ × ℕ)) (C : ℝ), 0 ≤ C ∧ ∀ (f : 𝓢(D, E)) (x : F), ‖x‖ ^ n.fst * ‖iteratedFDeriv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartzSeminormFamily 𝕜 D E) f) : 𝓢(D, E) →ₛₗ[σ] 𝓢(F, G) where @@ -700,7 +700,7 @@ For an example of using this definition, see `fderivCLM`. -/ def mkCLM [RingHomIsometric σ] (A : (D → E) → F → G) (hadd : ∀ (f g : 𝓢(D, E)) (x), A (f + g) x = A f x + A g x) (hsmul : ∀ (a : 𝕜) (f : 𝓢(D, E)) (x), A (a • f) x = σ a • A f x) - (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ⊤ (A f)) + (hsmooth : ∀ f : 𝓢(D, E), ContDiff ℝ ∞ (A f)) (hbound : ∀ n : ℕ × ℕ, ∃ (s : Finset (ℕ × ℕ)) (C : ℝ), 0 ≤ C ∧ ∀ (f : 𝓢(D, E)) (x : F), ‖x‖ ^ n.fst * ‖iteratedFDeriv ℝ n.snd (A f) x‖ ≤ C * s.sup (schwartzSeminormFamily 𝕜 D E) f) : 𝓢(D, E) →SL[σ] 𝓢(F, G) where @@ -740,19 +740,16 @@ variable [NormedField 𝕜] [NormedSpace 𝕜 F] [SMulCommClass ℝ 𝕜 F] /-- The map applying a vector to Hom-valued Schwartz function as a continuous linear map. -/ protected def evalCLM (m : E) : 𝓢(E, E →L[ℝ] F) →L[𝕜] 𝓢(E, F) := mkCLM (fun f x => f x m) (fun _ _ _ => rfl) (fun _ _ _ => rfl) - (fun f => ContDiff.clm_apply f.2 contDiff_const) - (by - rintro ⟨k, n⟩ - use {(k, n)}, ‖m‖, norm_nonneg _ - intro f x - refine - le_trans - (mul_le_mul_of_nonneg_left (norm_iteratedFDeriv_clm_apply_const f.2 le_top) - (by positivity)) - ?_ - move_mul [‖m‖] - gcongr ?_ * ‖m‖ - simp only [Finset.sup_singleton, schwartzSeminormFamily_apply, le_seminorm]) + (fun f => ContDiff.clm_apply f.2 contDiff_const) <| by + rintro ⟨k, n⟩ + use {(k, n)}, ‖m‖, norm_nonneg _ + intro f x + refine le_trans + (mul_le_mul_of_nonneg_left (norm_iteratedFDeriv_clm_apply_const f.2 (mod_cast le_top)) + (by positivity)) ?_ + move_mul [‖m‖] + gcongr ?_ * ‖m‖ + simp only [Finset.sup_singleton, schwartzSeminormFamily_apply, le_seminorm] end EvalCLM @@ -782,7 +779,8 @@ def bilinLeftCLM (B : E →L[ℝ] F →L[ℝ] G) {g : D → F} (hg : g.HasTemper intro f x have hxk : 0 ≤ ‖x‖ ^ k := by positivity have hnorm_mul := - ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) le_top + ContinuousLinearMap.norm_iteratedFDeriv_le_of_bilinear B f.smooth' hg.1 x (n := n) + (mod_cast le_top) refine le_trans (mul_le_mul_of_nonneg_left hnorm_mul hxk) ?_ move_mul [← ‖B‖] simp_rw [mul_assoc ‖B‖] @@ -864,7 +862,7 @@ def compCLM {g : D → E} (hg : g.HasTemperateGrowth) · exact le_trans (by simp [hC]) (le_self_pow₀ (by simp [hC]) hN₁') · refine le_self_pow₀ (one_le_pow₀ ?_) hN₁' simp only [le_add_iff_nonneg_right, norm_nonneg] - have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 le_top x hbound hgrowth' + have := norm_iteratedFDeriv_comp_le f.smooth' hg.1 (mod_cast le_top) x hbound hgrowth' have hxk : ‖x‖ ^ k ≤ (1 + ‖x‖) ^ k := pow_le_pow_left₀ (norm_nonneg _) (by simp only [zero_le_one, le_add_iff_nonneg_left]) _ refine le_trans (mul_le_mul hxk this (by positivity) (by positivity)) ?_ @@ -1018,7 +1016,8 @@ theorem iteratedPDeriv_eq_iteratedFDeriv {n : ℕ} {m : Fin n → E} {f : 𝓢(E simp only [iteratedPDeriv_succ_left, iteratedFDeriv_succ_apply_left] rw [← fderiv_continuousMultilinear_apply_const_apply] · simp [← ih] - · exact f.smooth'.differentiable_iteratedFDeriv (WithTop.coe_lt_top n) _ + · exact f.smooth'.differentiable_iteratedFDeriv (mod_cast ENat.coe_lt_top n) x + end Derivatives @@ -1053,7 +1052,7 @@ lemma integrable_pow_mul_iteratedFDeriv (f : 𝓢(D, V)) (k n : ℕ) : Integrable (fun x ↦ ‖x‖ ^ k * ‖iteratedFDeriv ℝ n f x‖) μ := integrable_of_le_of_pow_mul_le (norm_iteratedFDeriv_le_seminorm ℝ _ _) (le_seminorm ℝ _ _ _) - ((f.smooth ⊤).continuous_iteratedFDeriv le_top).aestronglyMeasurable + ((f.smooth ⊤).continuous_iteratedFDeriv (mod_cast le_top)).aestronglyMeasurable variable (μ) in lemma integrable_pow_mul (f : 𝓢(D, V)) diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 5b882dad07252..0151b54688137 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -81,7 +81,7 @@ noncomputable section open Real Complex MeasureTheory Filter TopologicalSpace -open scoped FourierTransform Topology +open scoped FourierTransform Topology ContDiff -- without this local instance, Lean tries first the instance -- `secondCountableTopologyEither_of_right` (whose priority is 100) and takes a very long time to @@ -310,7 +310,8 @@ lemma _root_.Continuous.fourierPowSMulRight {f : V → E} (hf : Continuous f) (n apply (smulRightL ℝ (fun (_ : Fin n) ↦ W) E).continuous₂.comp₂ _ hf exact Continuous.comp (map_continuous _) (continuous_pi (fun _ ↦ L.continuous)) -lemma _root_.ContDiff.fourierPowSMulRight {f : V → E} {k : ℕ∞} (hf : ContDiff ℝ k f) (n : ℕ) : +lemma _root_.ContDiff.fourierPowSMulRight + {f : V → E} {k : WithTop ℕ∞} (hf : ContDiff ℝ k f) (n : ℕ) : ContDiff ℝ k (fun v ↦ fourierPowSMulRight L f v n) := by simp_rw [fourierPowSMulRight_eq_comp] apply ContDiff.const_smul @@ -335,7 +336,7 @@ set_option maxSynthPendingDepth 2 in /-- The iterated derivative of a function multiplied by `(L v ⬝) ^ n` can be controlled in terms of the iterated derivatives of the initial function. -/ lemma norm_iteratedFDeriv_fourierPowSMulRight - {f : V → E} {K : ℕ∞} {C : ℝ} (hf : ContDiff ℝ K f) {n : ℕ} {k : ℕ} (hk : k ≤ K) + {f : V → E} {K : WithTop ℕ∞} {C : ℝ} (hf : ContDiff ℝ K f) {n : ℕ} {k : ℕ} (hk : k ≤ K) {v : V} (hv : ∀ i ≤ k, ∀ j ≤ n, ‖v‖ ^ j * ‖iteratedFDeriv ℝ i f v‖ ≤ C) : ‖iteratedFDeriv ℝ k (fun v ↦ fourierPowSMulRight L f v n) v‖ ≤ (2 * π) ^ n * (2 * n + 2) ^ k * ‖L‖ ^ n * C := by @@ -552,7 +553,8 @@ theorem fourierIntegral_iteratedFDeriv [FiniteDimensional ℝ V] ← fourierIntegral_continuousLinearMap_apply' J] exact H have h'n : n < N := (Nat.cast_lt.mpr n.lt_succ_self).trans_le hn - rw [fourierIntegral_fderiv _ (h'f n h'n.le) (hf.differentiable_iteratedFDeriv h'n) J] + rw [fourierIntegral_fderiv _ (h'f n h'n.le) + (hf.differentiable_iteratedFDeriv (mod_cast h'n)) J] simp only [ih h'n.le, fourierSMulRight_apply, ContinuousLinearMap.neg_apply, ContinuousLinearMap.flip_apply, neg_smul, smul_neg, neg_neg, smul_apply, fourierPowSMulRight_apply, ← coe_smul (E := E), smul_smul] @@ -583,9 +585,9 @@ theorem fourierPowSMulRight_iteratedFDeriv_fourierIntegral [FiniteDimensional simp only [Finset.mem_product, Finset.mem_range_succ_iff] at hp exact h'f _ _ ((Nat.cast_le.2 hp.1).trans hk) ((Nat.cast_le.2 hp.2).trans hm) apply (I.const_mul ((2 * π) ^ k * (2 * k + 2) ^ m * ‖L‖ ^ k)).mono' - ((hf.fourierPowSMulRight L k).continuous_iteratedFDeriv hm).aestronglyMeasurable + ((hf.fourierPowSMulRight L k).continuous_iteratedFDeriv (mod_cast hm)).aestronglyMeasurable filter_upwards with v - refine norm_iteratedFDeriv_fourierPowSMulRight _ hf hm (fun i hi j hj ↦ ?_) + refine norm_iteratedFDeriv_fourierPowSMulRight _ hf (mod_cast hm) (fun i hi j hj ↦ ?_) apply Finset.single_le_sum (f := fun p ↦ ‖v‖ ^ p.1 * ‖iteratedFDeriv ℝ p.2 f v‖) (a := (j, i)) · intro i _hi positivity @@ -616,7 +618,7 @@ theorem norm_fourierPowSMulRight_iteratedFDeriv_fourierIntegral_le [FiniteDimens · filter_upwards with v using norm_nonneg _ · exact (integrable_finset_sum _ I).const_mul _ · filter_upwards with v - apply norm_iteratedFDeriv_fourierPowSMulRight _ hf hn _ + apply norm_iteratedFDeriv_fourierPowSMulRight _ hf (mod_cast hn) _ intro i hi j hj apply Finset.single_le_sum (f := fun p ↦ ‖v‖ ^ p.1 * ‖iteratedFDeriv ℝ p.2 f v‖) (a := (j, i)) · intro i _hi diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 3a485d339830b..846339ecb2f00 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -90,7 +90,7 @@ theorem iter_deriv_exp : ∀ n : ℕ, deriv^[n] exp = exp | 0 => rfl | n + 1 => by rw [iterate_succ_apply, deriv_exp, iter_deriv_exp n] -theorem contDiff_exp {n : ℕ∞} : ContDiff 𝕜 n exp := +theorem contDiff_exp {n : WithTop ℕ∞} : ContDiff 𝕜 n exp := analyticOnNhd_cexp.restrictScalars.contDiff theorem hasStrictDerivAt_exp (x : ℂ) : HasStrictDerivAt exp (exp x) x := @@ -230,7 +230,7 @@ theorem hasStrictDerivAt_exp (x : ℝ) : HasStrictDerivAt exp (exp x) x := theorem hasDerivAt_exp (x : ℝ) : HasDerivAt exp (exp x) x := (Complex.hasDerivAt_exp x).real_of_complex -theorem contDiff_exp {n : ℕ∞} : ContDiff ℝ n exp := +theorem contDiff_exp {n : WithTop ℕ∞} : ContDiff ℝ n exp := Complex.contDiff_exp.real_of_complex @[simp] diff --git a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean index e9c7ca6362653..b886b2c966010 100644 --- a/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Log/Deriv.lean @@ -24,7 +24,7 @@ logarithm, derivative open Filter Finset Set -open scoped Topology +open scoped Topology ContDiff namespace Real @@ -66,12 +66,13 @@ theorem deriv_log (x : ℝ) : deriv log x = x⁻¹ := theorem deriv_log' : deriv log = Inv.inv := funext deriv_log -theorem contDiffOn_log {n : ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by - suffices ContDiffOn ℝ ⊤ log {0}ᶜ from this.of_le le_top +theorem contDiffOn_log {n : WithTop ℕ∞} : ContDiffOn ℝ n log {0}ᶜ := by + suffices ContDiffOn ℝ ω log {0}ᶜ from this.of_le le_top + rw [← contDiffOn_infty_iff_contDiffOn_omega] refine (contDiffOn_top_iff_deriv_of_isOpen isOpen_compl_singleton).2 ?_ simp [differentiableOn_log, contDiffOn_inv] -theorem contDiffAt_log {n : ℕ∞} : ContDiffAt ℝ n log x ↔ x ≠ 0 := +theorem contDiffAt_log {n : WithTop ℕ∞} : ContDiffAt ℝ n log x ↔ x ≠ 0 := ⟨fun h => continuousAt_log_iff.1 h.continuousAt, fun hx => (contDiffOn_log x hx).contDiffAt <| IsOpen.mem_nhds isOpen_compl_singleton hx⟩ diff --git a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean index 49c7cb8a86df9..0c619bd07f671 100644 --- a/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Pow/Deriv.lean @@ -282,7 +282,7 @@ theorem hasStrictFDerivAt_rpow_of_neg (p : ℝ × ℝ) (hp : p.1 < 0) : rw [div_eq_mul_inv, add_comm]; congr 2 <;> ring /-- The function `fun (x, y) => x ^ y` is infinitely smooth at `(x, y)` unless `x = 0`. -/ -theorem contDiffAt_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : ℕ∞} : +theorem contDiffAt_rpow_of_ne (p : ℝ × ℝ) (hp : p.1 ≠ 0) {n : WithTop ℕ∞} : ContDiffAt ℝ n (fun p : ℝ × ℝ => p.1 ^ p.2) p := by cases' hp.lt_or_lt with hneg hpos exacts @@ -358,7 +358,7 @@ theorem deriv_rpow_const' {p : ℝ} (h : 1 ≤ p) : (deriv fun x : ℝ => x ^ p) = fun x => p * x ^ (p - 1) := funext fun _ => deriv_rpow_const (Or.inr h) -theorem contDiffAt_rpow_const_of_ne {x p : ℝ} {n : ℕ∞} (h : x ≠ 0) : +theorem contDiffAt_rpow_const_of_ne {x p : ℝ} {n : WithTop ℕ∞} (h : x ≠ 0) : ContDiffAt ℝ n (fun x => x ^ p) x := (contDiffAt_rpow_of_ne (x, p) h).comp x (contDiffAt_id.prod contDiffAt_const) @@ -368,7 +368,8 @@ 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 [show ((n + 1 : ℕ) : ℕ∞) = n + 1 from rfl, contDiff_succ_iff_deriv, deriv_rpow_const' h1] + rw [show ((n + 1 : ℕ) : WithTop ℕ∞) = 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/Analysis/SpecialFunctions/SmoothTransition.lean b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean index 20b9a0ce94d94..a9815ec827dfa 100644 --- a/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean +++ b/Mathlib/Analysis/SpecialFunctions/SmoothTransition.lean @@ -113,7 +113,7 @@ theorem contDiff_polynomial_eval_inv_mul {n : ℕ∞} (p : ℝ[X]) : exact (hasDerivAt_polynomial_eval_inv_mul p _).deriv /-- The function `expNegInvGlue` is smooth. -/ -protected theorem contDiff {n} : ContDiff ℝ n expNegInvGlue := by +protected theorem contDiff {n : ℕ∞} : ContDiff ℝ n expNegInvGlue := by simpa using contDiff_polynomial_eval_inv_mul 1 end expNegInvGlue @@ -174,12 +174,12 @@ theorem lt_one_of_lt_one (h : x < 1) : smoothTransition x < 1 := theorem pos_of_pos (h : 0 < x) : 0 < smoothTransition x := div_pos (expNegInvGlue.pos_of_pos h) (pos_denom x) -protected theorem contDiff {n} : ContDiff ℝ n smoothTransition := +protected theorem contDiff {n : ℕ∞} : ContDiff ℝ n smoothTransition := expNegInvGlue.contDiff.div (expNegInvGlue.contDiff.add <| expNegInvGlue.contDiff.comp <| contDiff_const.sub contDiff_id) fun x => (pos_denom x).ne' -protected theorem contDiffAt {x n} : ContDiffAt ℝ n smoothTransition x := +protected theorem contDiffAt {x : ℝ} {n : ℕ∞} : ContDiffAt ℝ n smoothTransition x := smoothTransition.contDiff.contDiffAt protected theorem continuous : Continuous smoothTransition := diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean index e414b1421381a..a06b4a537e039 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/ComplexDeriv.lean @@ -62,7 +62,7 @@ theorem deriv_tan (x : ℂ) : deriv tan x = 1 / cos x ^ 2 := else (hasDerivAt_tan h).deriv @[simp] -theorem contDiffAt_tan {x : ℂ} {n : ℕ∞} : ContDiffAt ℂ n tan x ↔ cos x ≠ 0 := +theorem contDiffAt_tan {x : ℂ} {n : WithTop ℕ∞} : ContDiffAt ℂ n tan x ↔ cos x ≠ 0 := ⟨fun h => continuousAt_tan.1 h.continuousAt, contDiff_sin.contDiffAt.div contDiff_cos.contDiffAt⟩ end Complex diff --git a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean index 6ade7f02ca336..d9efa5b942979 100644 --- a/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/Trigonometric/InverseDeriv.lean @@ -14,7 +14,7 @@ Derivatives of `arcsin` and `arccos`. noncomputable section -open scoped Topology Filter Real +open scoped Topology Filter Real ContDiff open Set namespace Real @@ -22,7 +22,7 @@ namespace Real section Arcsin theorem deriv_arcsin_aux {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : - HasStrictDerivAt arcsin (1 / √(1 - x ^ 2)) x ∧ ContDiffAt ℝ ⊤ arcsin x := by + HasStrictDerivAt arcsin (1 / √(1 - x ^ 2)) x ∧ ContDiffAt ℝ ω arcsin x := by cases' h₁.lt_or_lt with h₁ h₁ · have : 1 - x ^ 2 < 0 := by nlinarith [h₁] rw [sqrt_eq_zero'.2 this.le, div_zero] @@ -50,7 +50,8 @@ theorem hasDerivAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : HasDerivAt arcsin (1 / √(1 - x ^ 2)) x := (hasStrictDerivAt_arcsin h₁ h₂).hasDerivAt -theorem contDiffAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : ContDiffAt ℝ n arcsin x := +theorem contDiffAt_arcsin {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : WithTop ℕ∞} : + ContDiffAt ℝ n arcsin x := (deriv_arcsin_aux h₁ h₂).2.of_le le_top theorem hasDerivWithinAt_arcsin_Ici {x : ℝ} (h : x ≠ -1) : @@ -102,12 +103,13 @@ theorem differentiableOn_arcsin : DifferentiableOn ℝ arcsin {-1, 1}ᶜ := fun (differentiableAt_arcsin.2 ⟨fun h => hx (Or.inl h), fun h => hx (Or.inr h)⟩).differentiableWithinAt -theorem contDiffOn_arcsin {n : ℕ∞} : ContDiffOn ℝ n arcsin {-1, 1}ᶜ := fun _x hx => +theorem contDiffOn_arcsin {n : WithTop ℕ∞} : ContDiffOn ℝ n arcsin {-1, 1}ᶜ := fun _x hx => (contDiffAt_arcsin (mt Or.inl hx) (mt Or.inr hx)).contDiffWithinAt -theorem contDiffAt_arcsin_iff {x : ℝ} {n : ℕ∞} : ContDiffAt ℝ n arcsin x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := +theorem contDiffAt_arcsin_iff {x : ℝ} {n : WithTop ℕ∞} : + ContDiffAt ℝ n arcsin x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := ⟨fun h => or_iff_not_imp_left.2 fun hn => differentiableAt_arcsin.1 <| h.differentiableAt <| - ENat.one_le_iff_ne_zero.2 hn, + ENat.one_le_iff_ne_zero_withTop.mpr hn, fun h => h.elim (fun hn => hn.symm ▸ (contDiff_zero.2 continuous_arcsin).contDiffAt) fun hx => contDiffAt_arcsin hx.1 hx.2⟩ @@ -123,7 +125,8 @@ theorem hasDerivAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) : HasDerivAt arccos (-(1 / √(1 - x ^ 2))) x := (hasDerivAt_arcsin h₁ h₂).const_sub (π / 2) -theorem contDiffAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : ℕ∞} : ContDiffAt ℝ n arccos x := +theorem contDiffAt_arccos {x : ℝ} (h₁ : x ≠ -1) (h₂ : x ≠ 1) {n : WithTop ℕ∞} : + ContDiffAt ℝ n arccos x := contDiffAt_const.sub (contDiffAt_arcsin h₁ h₂) theorem hasDerivWithinAt_arccos_Ici {x : ℝ} (h : x ≠ -1) : @@ -152,10 +155,10 @@ theorem deriv_arccos : deriv arccos = fun x => -(1 / √(1 - x ^ 2)) := theorem differentiableOn_arccos : DifferentiableOn ℝ arccos {-1, 1}ᶜ := differentiableOn_arcsin.const_sub _ -theorem contDiffOn_arccos {n : ℕ∞} : ContDiffOn ℝ n arccos {-1, 1}ᶜ := +theorem contDiffOn_arccos {n : WithTop ℕ∞} : ContDiffOn ℝ n arccos {-1, 1}ᶜ := contDiffOn_const.sub contDiffOn_arcsin -theorem contDiffAt_arccos_iff {x : ℝ} {n : ℕ∞} : +theorem contDiffAt_arccos_iff {x : ℝ} {n : WithTop ℕ∞} : ContDiffAt ℝ n arccos x ↔ n = 0 ∨ x ≠ -1 ∧ x ≠ 1 := by refine Iff.trans ⟨fun h => ?_, fun h => ?_⟩ contDiffAt_arcsin_iff <;> simpa [arccos] using (contDiffAt_const (c := π / 2)).sub h diff --git a/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean b/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean index 998cc71804f14..9e2358ae486b1 100644 --- a/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean +++ b/Mathlib/Geometry/Manifold/Algebra/LeftInvariantDerivation.lean @@ -23,6 +23,9 @@ implementing one of the possible definitions of the Lie algebra attached to a Li noncomputable section open scoped LieGroup Manifold Derivation +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (G : Type*) diff --git a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean index 3808ac40a51c1..7e6a90b613a22 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Monoid.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Monoid.lean @@ -18,6 +18,9 @@ semigroups. open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) library_note "Design choices about smooth algebraic structures"/-- 1. All smooth algebraic structures on `G` are `Prop`-valued classes that extend diff --git a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean index e8adb09c03d94..db675c91210d4 100644 --- a/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean +++ b/Mathlib/Geometry/Manifold/Algebra/SmoothFunctions.lean @@ -15,6 +15,9 @@ In this file, we define instances of algebraic structures over smooth functions. noncomputable section open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) open TopologicalSpace diff --git a/Mathlib/Geometry/Manifold/Algebra/Structures.lean b/Mathlib/Geometry/Manifold/Algebra/Structures.lean index da593ec373f61..a20f6b529acf8 100644 --- a/Mathlib/Geometry/Manifold/Algebra/Structures.lean +++ b/Mathlib/Geometry/Manifold/Algebra/Structures.lean @@ -12,7 +12,7 @@ In this file we define smooth structures that build on Lie groups. We prefer usi instead of Lie mainly because Lie ring has currently another use in mathematics. -/ -open scoped Manifold +open scoped Manifold ContDiff section SmoothRing @@ -35,7 +35,7 @@ instance (priority := 100) SmoothRing.toSmoothMul (I : ModelWithCorners 𝕜 E H -- see Note [lower instance priority] instance (priority := 100) SmoothRing.toLieAddGroup (I : ModelWithCorners 𝕜 E H) (R : Type*) [Ring R] [TopologicalSpace R] [ChartedSpace H R] [SmoothRing I R] : LieAddGroup I R where - compatible := StructureGroupoid.compatible (contDiffGroupoid ⊤ I) + compatible := StructureGroupoid.compatible (contDiffGroupoid ∞ I) smooth_add := contMDiff_add I smooth_neg := by simpa only [neg_one_mul] using contMDiff_mul_left (G := R) (a := -1) diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean index 2036cda540d21..31ff5e240b2d7 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Atlas.lean @@ -13,7 +13,7 @@ and that local structomorphisms are smooth with smooth inverses. -/ open Set ChartedSpace SmoothManifoldWithCorners -open scoped Manifold +open scoped Manifold ContDiff variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] -- declare a smooth manifold `M` over the pair `(E, H)`. @@ -46,18 +46,14 @@ theorem contMDiffOn_model_symm : ContMDiffOn 𝓘(𝕜, E) I n I.symm (range I) /-- An atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e e.source := - ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_of_mem_maximalAtlas - contDiffWithinAtProp_id h) - le_top + ContMDiffOn.of_le ((contDiffWithinAt_localInvariantProp ⊤).liftPropOn_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top /-- The inverse of an atlas member is `C^n` for any `n`. -/ theorem contMDiffOn_symm_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) : ContMDiffOn I I n e.symm e.target := - ContMDiffOn.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftPropOn_symm_of_mem_maximalAtlas - contDiffWithinAtProp_id h) - le_top + ContMDiffOn.of_le ((contDiffWithinAt_localInvariantProp ⊤).liftPropOn_symm_of_mem_maximalAtlas + contDiffWithinAtProp_id h) le_top theorem contMDiffAt_of_mem_maximalAtlas (h : e ∈ maximalAtlas I M) (hx : x ∈ e.source) : ContMDiffAt I I n e x := @@ -112,7 +108,7 @@ theorem contMDiffWithinAt_extChartAt_symm_range /-- An element of `contDiffGroupoid ⊤ I` is `C^n` for any `n`. -/ theorem contMDiffOn_of_mem_contDiffGroupoid {e' : PartialHomeomorph H H} - (h : e' ∈ contDiffGroupoid ⊤ I) : ContMDiffOn I I n e' e'.source := + (h : e' ∈ contDiffGroupoid ∞ I) : ContMDiffOn I I n e' e'.source := (contDiffWithinAt_localInvariantProp n).liftPropOn_of_mem_groupoid contDiffWithinAtProp_id h end Atlas @@ -124,7 +120,7 @@ section IsLocalStructomorph variable [ChartedSpace H M'] [IsM' : SmoothManifoldWithCorners I M'] theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M M'} - (hf : LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source) : + (hf : LiftPropOn (contDiffGroupoid ∞ I).IsLocalStructomorphWithinAt f f.source) : ContMDiffOn I I ⊤ f f.source := by -- It suffices to show smoothness near each `x` apply contMDiffOn_of_locally_contMDiffOn @@ -170,7 +166,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff_aux {f : PartialHomeomorph M is a local structomorphism for `I`, if and only if it is manifold-smooth on the domain of definition in both directions. -/ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') : - LiftPropOn (contDiffGroupoid ⊤ I).IsLocalStructomorphWithinAt f f.source ↔ + LiftPropOn (contDiffGroupoid ∞ I).IsLocalStructomorphWithinAt f f.source ↔ ContMDiffOn I I ⊤ f f.source ∧ ContMDiffOn I I ⊤ f.symm f.target := by constructor · intro h @@ -186,7 +182,7 @@ theorem isLocalStructomorphOn_contDiffGroupoid_iff (f : PartialHomeomorph M M') refine ⟨(f.symm.continuousAt hX).continuousWithinAt, fun h2x => ?_⟩ obtain ⟨e, he, h2e, hef, hex⟩ : ∃ e : PartialHomeomorph H H, - e ∈ contDiffGroupoid ⊤ I ∧ + e ∈ contDiffGroupoid ∞ I ∧ e.source ⊆ (c.symm ≫ₕ f ≫ₕ c').source ∧ EqOn (c' ∘ f ∘ c.symm) e e.source ∧ c x ∈ e.source := by have h1 : c' = chartAt H (f x) := by simp only [f.right_inv hX] diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean index f5506bc51360d..bbe1eedcdb768 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Basic.lean @@ -156,7 +156,7 @@ section id theorem contMDiff_id : ContMDiff I I n (id : M → M) := ContMDiff.of_le - ((contDiffWithinAt_localInvariantProp ∞).liftProp_id contDiffWithinAtProp_id) le_top + ((contDiffWithinAt_localInvariantProp ⊤).liftProp_id contDiffWithinAtProp_id) le_top @[deprecated (since := "2024-11-20")] alias smooth_id := contMDiff_id diff --git a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean index aa5e8d957cfd7..49996ee7ba8cb 100644 --- a/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean +++ b/Mathlib/Geometry/Manifold/ContMDiff/Defs.lean @@ -46,7 +46,7 @@ in terms of extended charts in `contMDiffOn_iff` and `contMDiff_iff`. open Set Function Filter ChartedSpace SmoothManifoldWithCorners -open scoped Topology Manifold +open scoped Topology Manifold ContDiff /-! ### Definition of smooth functions between manifolds -/ @@ -110,7 +110,8 @@ theorem contDiffWithinAt_localInvariantProp (n : ℕ∞) : rw [this] at h have : I (e x) ∈ I.symm ⁻¹' e.target ∩ range I := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he).2.contDiffWithinAt this - convert (h.comp_inter _ (this.of_le le_top)).mono_of_mem_nhdsWithin _ using 1 + convert (h.comp_inter _ (this.of_le (mod_cast le_top))).mono_of_mem_nhdsWithin _ + using 1 · ext y; simp only [mfld_simps] refine mem_nhdsWithin.mpr ⟨I.symm ⁻¹' e.target, e.open_target.preimage I.continuous_symm, by @@ -127,7 +128,7 @@ theorem contDiffWithinAt_localInvariantProp (n : ℕ∞) : have A : (I' ∘ f ∘ I.symm) (I x) ∈ I'.symm ⁻¹' e'.source ∩ range I' := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he').1.contDiffWithinAt A - convert (this.of_le le_top).comp _ h _ + convert (this.of_le (mod_cast le_top)).comp _ h _ · ext y; simp only [mfld_simps] · intro y hy; simp only [mfld_simps] at hy; simpa only [hy, mfld_simps] using hs hy.1 @@ -197,7 +198,7 @@ def ContMDiff (n : ℕ∞) (f : M → M') := theorem ContMDiffWithinAt.of_le (hf : ContMDiffWithinAt I I' n f s x) (le : m ≤ n) : ContMDiffWithinAt I I' m f s x := by simp only [ContMDiffWithinAt, LiftPropWithinAt] at hf ⊢ - exact ⟨hf.1, hf.2.of_le le⟩ + exact ⟨hf.1, hf.2.of_le (mod_cast le)⟩ theorem ContMDiffAt.of_le (hf : ContMDiffAt I I' n f x) (le : m ≤ n) : ContMDiffAt I I' m f x := ContMDiffWithinAt.of_le hf le diff --git a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean index ad3c237062762..38b8fd8f39f73 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMFDeriv.lean @@ -113,7 +113,7 @@ protected theorem ContMDiffWithinAt.mfderivWithin {x₀ : N} {f : N → M → M' rw [contMDiffWithinAt_iff] at hf' hg' simp_rw [Function.comp_def, uncurry, extChartAt_prod, PartialEquiv.prod_coe_symm, ModelWithCorners.range_prod] at hf' ⊢ - apply ContDiffWithinAt.fderivWithin _ _ _ hmn + apply ContDiffWithinAt.fderivWithin _ _ _ (show (m : WithTop ℕ∞) + 1 ≤ n from mod_cast hmn ) · simp [hx₀, t'] · apply inter_subset_left.trans rw [preimage_subset_iff] @@ -404,6 +404,9 @@ namespace ContMDiffMap -- (However as a consequence we import `Mathlib/Geometry/Manifold/ContMDiffMap.lean` here now.) -- They could be moved to another file (perhaps a new file) if desired. open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) protected theorem mdifferentiable' (f : C^n⟮I, M; I', M'⟯) (hn : 1 ≤ n) : MDifferentiable I I' f := f.contMDiff.mdifferentiable hn diff --git a/Mathlib/Geometry/Manifold/ContMDiffMap.lean b/Mathlib/Geometry/Manifold/ContMDiffMap.lean index f8bd9f7980ad8..2be46a551a705 100644 --- a/Mathlib/Geometry/Manifold/ContMDiffMap.lean +++ b/Mathlib/Geometry/Manifold/ContMDiffMap.lean @@ -39,6 +39,10 @@ scoped[Manifold] notation "C^" n "⟮" I ", " M "; " k "⟯" => ContMDiffMap I (modelWithCornersSelf k k) M k n open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + namespace ContMDiffMap diff --git a/Mathlib/Geometry/Manifold/DerivationBundle.lean b/Mathlib/Geometry/Manifold/DerivationBundle.lean index ca029181286c4..fb3dbd766186a 100644 --- a/Mathlib/Geometry/Manifold/DerivationBundle.lean +++ b/Mathlib/Geometry/Manifold/DerivationBundle.lean @@ -25,6 +25,9 @@ variable (𝕜 : Type*) [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [TopologicalSpace M] [ChartedSpace H M] (n : ℕ∞) open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) -- the following two instances prevent poorly understood type class inference timeout problems instance smoothFunctionsAlgebra : Algebra 𝕜 C^∞⟮I, M; 𝕜⟯ := by infer_instance diff --git a/Mathlib/Geometry/Manifold/Diffeomorph.lean b/Mathlib/Geometry/Manifold/Diffeomorph.lean index 5ee61a8605fe2..6a6dd316cf875 100644 --- a/Mathlib/Geometry/Manifold/Diffeomorph.lean +++ b/Mathlib/Geometry/Manifold/Diffeomorph.lean @@ -44,7 +44,7 @@ diffeomorphism, manifold -/ -open scoped Manifold Topology +open scoped Manifold Topology ContDiff open Function Set @@ -491,7 +491,7 @@ instance smoothManifoldWithCorners_transDiffeomorph [SmoothManifoldWithCorners I SmoothManifoldWithCorners (I.transDiffeomorph e) M := by refine smoothManifoldWithCorners_of_contDiffOn (I.transDiffeomorph e) M fun e₁ e₂ h₁ h₂ => ?_ refine e.contDiff.comp_contDiffOn - (((contDiffGroupoid ⊤ I).compatible h₁ h₂).1.comp e.symm.contDiff.contDiffOn ?_) + (((contDiffGroupoid ∞ I).compatible h₁ h₂).1.comp e.symm.contDiff.contDiffOn ?_) simp only [mapsTo_iff_subset_preimage] mfld_set_tac diff --git a/Mathlib/Geometry/Manifold/Instances/Real.lean b/Mathlib/Geometry/Manifold/Instances/Real.lean index 26caf2f7bee47..ce5fb5a376760 100644 --- a/Mathlib/Geometry/Manifold/Instances/Real.lean +++ b/Mathlib/Geometry/Manifold/Instances/Real.lean @@ -42,7 +42,7 @@ noncomputable section open Set Function -open scoped Manifold +open scoped Manifold ContDiff /-- The half-space in `ℝ^n`, used to model manifolds with boundary. We only define it when `1 ≤ n`, as the definition only makes sense in this case. diff --git a/Mathlib/Geometry/Manifold/Instances/Sphere.lean b/Mathlib/Geometry/Manifold/Instances/Sphere.lean index 00100a2fdd4ed..a426774e69c58 100644 --- a/Mathlib/Geometry/Manifold/Instances/Sphere.lean +++ b/Mathlib/Geometry/Manifold/Instances/Sphere.lean @@ -69,7 +69,7 @@ noncomputable section open Metric Module Function -open scoped Manifold +open scoped Manifold ContDiff section StereographicProjection @@ -93,7 +93,7 @@ theorem stereoToFun_apply (x : E) : rfl theorem contDiffOn_stereoToFun : - ContDiffOn ℝ ⊤ (stereoToFun v) {x : E | innerSL _ v x ≠ (1 : ℝ)} := by + ContDiffOn ℝ ∞ (stereoToFun v) {x : E | innerSL _ v x ≠ (1 : ℝ)} := by refine ContDiffOn.smul ?_ (orthogonalProjection (ℝ ∙ v)ᗮ).contDiff.contDiffOn refine contDiff_const.contDiffOn.div ?_ ?_ · exact (contDiff_const.sub (innerSL ℝ v).contDiff).contDiffOn @@ -157,13 +157,13 @@ theorem hasFDerivAt_stereoInvFunAux_comp_coe (v : E) : hasFDerivAt_stereoInvFunAux v refine this.comp (0 : (ℝ ∙ v)ᗮ) (by apply ContinuousLinearMap.hasFDerivAt) -theorem contDiff_stereoInvFunAux : ContDiff ℝ ⊤ (stereoInvFunAux v) := by - have h₀ : ContDiff ℝ ⊤ fun w : E => ‖w‖ ^ 2 := contDiff_norm_sq ℝ - have h₁ : ContDiff ℝ ⊤ fun w : E => (‖w‖ ^ 2 + 4)⁻¹ := by +theorem contDiff_stereoInvFunAux : ContDiff ℝ ∞ (stereoInvFunAux v) := by + have h₀ : ContDiff ℝ ∞ fun w : E => ‖w‖ ^ 2 := contDiff_norm_sq ℝ + have h₁ : ContDiff ℝ ∞ fun w : E => (‖w‖ ^ 2 + 4)⁻¹ := by refine (h₀.add contDiff_const).inv ?_ intro x - positivity - have h₂ : ContDiff ℝ ⊤ fun w => (4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v := by + nlinarith + have h₂ : ContDiff ℝ ∞ fun w => (4 : ℝ) • w + (‖w‖ ^ 2 - 4) • v := by refine (contDiff_const.smul contDiff_id).add ?_ exact (h₀.sub contDiff_const).smul contDiff_const exact h₁.smul h₂ @@ -389,7 +389,7 @@ instance EuclideanSpace.instSmoothManifoldWithCornersSphere {n : ℕ} [Fact (fin -- Porting note: need to help with implicit variables again have H₂ := (contDiff_stereoInvFunAux (v := v.val)|>.comp (ℝ ∙ (v : E))ᗮ.subtypeL.contDiff).comp U.symm.contDiff - convert H₁.comp_inter (H₂.contDiffOn : ContDiffOn ℝ ⊤ _ Set.univ) using 1 + convert H₁.comp_inter (H₂.contDiffOn : ContDiffOn ℝ ∞ _ Set.univ) using 1 -- -- squeezed from `ext, simp [sphere_ext_iff, stereographic'_symm_apply, real_inner_comm]` simp only [PartialHomeomorph.trans_toPartialEquiv, PartialHomeomorph.symm_toPartialEquiv, PartialEquiv.trans_source, PartialEquiv.symm_source, stereographic'_target, @@ -407,7 +407,7 @@ instance (n : ℕ) : /-- The inclusion map (i.e., `coe`) from the sphere in `E` to `E` is smooth. -/ theorem contMDiff_coe_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] : - ContMDiff (𝓡 n) 𝓘(ℝ, E) ∞ ((↑) : sphere (0 : E) 1 → E) := by + ContMDiff (𝓡 n) 𝓘(ℝ, E) ⊤ ((↑) : sphere (0 : E) 1 → E) := by -- Porting note: trouble with filling these implicit variables in the instance have := EuclideanSpace.instSmoothManifoldWithCornersSphere (E := E) (n := n) rw [contMDiff_iff] @@ -438,7 +438,7 @@ theorem ContMDiff.codRestrict_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] {m (-- Again, partially removing type ascription... Weird that this helps! OrthonormalBasis.fromOrthogonalSpanSingleton n (ne_zero_of_mem_unit_sphere (-v))).repr - have h : ContDiffOn ℝ ⊤ _ Set.univ := U.contDiff.contDiffOn + have h : ContDiffOn ℝ ∞ _ Set.univ := U.contDiff.contDiffOn have H₁ := (h.comp_inter contDiffOn_stereoToFun).contMDiffOn have H₂ : ContMDiffOn _ _ _ _ Set.univ := hf.contMDiffOn convert (H₁.of_le le_top).comp' H₂ using 1 @@ -453,7 +453,7 @@ theorem ContMDiff.codRestrict_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] {m /-- The antipodal map is smooth. -/ theorem contMDiff_neg_sphere {n : ℕ} [Fact (finrank ℝ E = n + 1)] : - ContMDiff (𝓡 n) (𝓡 n) ∞ fun x : sphere (0 : E) 1 => -x := by + ContMDiff (𝓡 n) (𝓡 n) ⊤ fun x : sphere (0 : E) 1 => -x := by -- this doesn't elaborate well in term mode apply ContMDiff.codRestrict_sphere apply contDiff_neg.contMDiff.comp _ @@ -555,7 +555,7 @@ instance : LieGroup (𝓡 1) Circle where smooth_mul := by apply ContMDiff.codRestrict_sphere let c : Circle → ℂ := (↑) - have h₂ : ContMDiff (𝓘(ℝ, ℂ).prod 𝓘(ℝ, ℂ)) 𝓘(ℝ, ℂ) ∞ fun z : ℂ × ℂ => z.fst * z.snd := by + have h₂ : ContMDiff (𝓘(ℝ, ℂ).prod 𝓘(ℝ, ℂ)) 𝓘(ℝ, ℂ) ⊤ fun z : ℂ × ℂ => z.fst * z.snd := by rw [contMDiff_iff] exact ⟨continuous_mul, fun x y => contDiff_mul.contDiffOn⟩ -- Porting note: needed to fill in first 3 arguments or could not figure out typeclasses @@ -569,7 +569,7 @@ instance : LieGroup (𝓡 1) Circle where exact Complex.conjCLE.contDiff.contMDiff.comp contMDiff_coe_sphere /-- The map `fun t ↦ exp (t * I)` from `ℝ` to the unit circle in `ℂ` is smooth. -/ -theorem contMDiff_circleExp : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ∞ Circle.exp := +theorem contMDiff_circleExp : ContMDiff 𝓘(ℝ, ℝ) (𝓡 1) ⊤ Circle.exp := (contDiff_exp.comp (contDiff_id.smul contDiff_const)).contMDiff.codRestrict_sphere _ @[deprecated (since := "2024-07-25")] alias contMDiff_expMapCircle := contMDiff_circleExp diff --git a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean index 8ee99e793fb7c..b66b8302d348a 100644 --- a/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean +++ b/Mathlib/Geometry/Manifold/Instances/UnitsOfNormedAlgebra.lean @@ -27,6 +27,9 @@ example {V : Type*} [NormedAddCommGroup V] [NormedSpace 𝕜 V] [CompleteSpace V noncomputable section open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) namespace Units diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean index eb88f2a6c5c87..9386f8ad261ec 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Atlas.lean @@ -27,7 +27,7 @@ charts, differentiable, bijective noncomputable section -open scoped Manifold +open scoped Manifold ContDiff open Bundle Set Topology variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] @@ -99,7 +99,7 @@ theorem mdifferentiableAt_atlas (h : e ∈ atlas H M) {x : M} (hx : x ∈ e.sour ContDiffOn 𝕜 ∞ (I ∘ (chartAt H x).symm.trans e ∘ I.symm) (I.symm ⁻¹' ((chartAt H x).symm.trans e).source ∩ range I) := this.1 - have B := A.differentiableOn le_top (I ((chartAt H x : M → H) x)) mem + have B := A.differentiableOn (mod_cast le_top) (I ((chartAt H x : M → H) x)) mem simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] @@ -120,7 +120,7 @@ theorem mdifferentiableAt_atlas_symm (h : e ∈ atlas H M) {x : H} (hx : x ∈ e ContDiffOn 𝕜 ∞ (I ∘ e.symm.trans (chartAt H (e.symm x)) ∘ I.symm) (I.symm ⁻¹' (e.symm.trans (chartAt H (e.symm x))).source ∩ range I) := this.1 - have B := A.differentiableOn le_top (I x) mem + have B := A.differentiableOn (mod_cast le_top) (I x) mem simp only [mfld_simps] at B rw [inter_comm, differentiableWithinAt_inter] at B · simpa only [mfld_simps] diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean index 377ab9d17b795..acf6be63394ec 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Basic.lean @@ -467,7 +467,8 @@ theorem ContMDiffWithinAt.mdifferentiableWithinAt (hf : ContMDiffWithinAt I I' n apply hf.1.preimage_mem_nhdsWithin exact extChartAt_source_mem_nhds (f x) rw [mdifferentiableWithinAt_iff] - exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt hn).mono (by mfld_set_tac)⟩ + exact ⟨hf.1.mono inter_subset_left, (hf.2.differentiableWithinAt (mod_cast hn)).mono + (by mfld_set_tac)⟩ theorem ContMDiffAt.mdifferentiableAt (hf : ContMDiffAt I I' n f x) (hn : 1 ≤ n) : MDifferentiableAt I I' f x := diff --git a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean index 79091c4619038..b7c3206ec3989 100644 --- a/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean +++ b/Mathlib/Geometry/Manifold/MFDeriv/Defs.lean @@ -99,7 +99,7 @@ derivative, manifold noncomputable section -open scoped Topology +open scoped Topology ContDiff open Set ChartedSpace section DerivativesDefinitions @@ -147,7 +147,7 @@ theorem differentiableWithinAtProp_self_target {f : H → E'} {s : Set H} {x : H /-- Being differentiable in the model space is a local property, invariant under smooth maps. Therefore, it will lift nicely to manifolds. -/ theorem differentiableWithinAt_localInvariantProp : - (contDiffGroupoid ⊤ I).LocalInvariantProp (contDiffGroupoid ⊤ I') + (contDiffGroupoid ∞ I).LocalInvariantProp (contDiffGroupoid ∞ I') (DifferentiableWithinAtProp I I') := { is_local := by intro s x u f u_open xu @@ -167,7 +167,8 @@ theorem differentiableWithinAt_localInvariantProp : rw [this] at h have : I (e x) ∈ I.symm ⁻¹' e.target ∩ Set.range I := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he).2.contDiffWithinAt this - convert (h.comp' _ (this.differentiableWithinAt le_top)).mono_of_mem_nhdsWithin _ using 1 + convert (h.comp' _ (this.differentiableWithinAt (mod_cast le_top))).mono_of_mem_nhdsWithin _ + using 1 · ext y; simp only [mfld_simps] refine mem_nhdsWithin.mpr @@ -187,7 +188,7 @@ theorem differentiableWithinAt_localInvariantProp : have A : (I' ∘ f ∘ I.symm) (I x) ∈ I'.symm ⁻¹' e'.source ∩ Set.range I' := by simp only [hx, mfld_simps] have := (mem_groupoid_of_pregroupoid.2 he').1.contDiffWithinAt A - convert (this.differentiableWithinAt le_top).comp _ h _ + convert (this.differentiableWithinAt (mod_cast le_top)).comp _ h _ · ext y; simp only [mfld_simps] · intro y hy; simp only [mfld_simps] at hy; simpa only [hy, mfld_simps] using hs hy.1 } diff --git a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean index 3d001e4e01c5d..5a4131d403b71 100644 --- a/Mathlib/Geometry/Manifold/PartitionOfUnity.lean +++ b/Mathlib/Geometry/Manifold/PartitionOfUnity.lean @@ -59,6 +59,9 @@ universe uι uE uH uM uF open Function Filter Module Set open scoped Topology Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) noncomputable section diff --git a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean index 48aec60a959fa..5c91ede6dc259 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/LocallyRingedSpace.lean @@ -31,6 +31,10 @@ smooth manifolds. noncomputable section universe u +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + variable {𝕜 : Type u} [NontriviallyNormedField 𝕜] [CompleteSpace 𝕜] {EM : Type*} [NormedAddCommGroup EM] [NormedSpace 𝕜 EM] {HM : Type*} [TopologicalSpace HM] (IM : ModelWithCorners 𝕜 EM HM) diff --git a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean index 0f528359115a1..d794651ff25d8 100644 --- a/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean +++ b/Mathlib/Geometry/Manifold/Sheaf/Smooth.lean @@ -66,6 +66,10 @@ https://github.com/leanprover-community/mathlib4/pull/5726. noncomputable section open TopologicalSpace Opposite +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) + universe u variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index df9b1d40d1761..1e7be0aaf95d1 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -133,10 +133,7 @@ universe u v w u' v' w' open Set Filter Function -open scoped Manifold Filter Topology - -/-- The extended natural number `∞` -/ -scoped[Manifold] notation "∞" => (⊤ : ℕ∞) +open scoped Manifold Filter Topology ContDiff /-! ### Models with corners. -/ @@ -529,9 +526,9 @@ section contDiffGroupoid /-! ### Smooth functions on models with corners -/ -variable {m n : ℕ∞} {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] - [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] {I : ModelWithCorners 𝕜 E H} {M : Type*} - [TopologicalSpace M] +variable {m n : WithTop ℕ∞} {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} + [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] + {I : ModelWithCorners 𝕜 E H} {M : Type*} [TopologicalSpace M] variable (n I) in /-- Given a model with corners `(E, H)`, we define the pregroupoid of `C^n` transformations of `H` @@ -623,8 +620,8 @@ variable {E' H' : Type*} [NormedAddCommGroup E'] [NormedSpace 𝕜 E'] [Topologi /-- The product of two smooth partial homeomorphisms is smooth. -/ theorem contDiffGroupoid_prod {I : ModelWithCorners 𝕜 E H} {I' : ModelWithCorners 𝕜 E' H'} - {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid ⊤ I) - (he' : e' ∈ contDiffGroupoid ⊤ I') : e.prod e' ∈ contDiffGroupoid ⊤ (I.prod I') := by + {e : PartialHomeomorph H H} {e' : PartialHomeomorph H' H'} (he : e ∈ contDiffGroupoid ∞ I) + (he' : e' ∈ contDiffGroupoid ∞ I') : e.prod e' ∈ contDiffGroupoid ∞ (I.prod I') := by cases' he with he he_symm cases' he' with he' he'_symm simp only at he he_symm he' he'_symm @@ -672,7 +669,7 @@ theorem smoothManifoldWithCorners_of_contDiffOn {𝕜 : Type*} [NontriviallyNorm {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) (M : Type*) [TopologicalSpace M] [ChartedSpace H M] (h : ∀ e e' : PartialHomeomorph M H, e ∈ atlas H M → e' ∈ atlas H M → - ContDiffOn 𝕜 ⊤ (I ∘ e.symm ≫ₕ e' ∘ I.symm) (I.symm ⁻¹' (e.symm ≫ₕ e').source ∩ range I)) : + ContDiffOn 𝕜 ∞ (I ∘ e.symm ≫ₕ e' ∘ I.symm) (I.symm ⁻¹' (e.symm ≫ₕ e').source ∩ range I)) : SmoothManifoldWithCorners I M where compatible := by haveI : HasGroupoid M (contDiffGroupoid ∞ I) := hasGroupoid_of_pregroupoid _ (h _ _) @@ -741,8 +738,8 @@ instance prod {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedA compatible := by rintro f g ⟨f1, hf1, f2, hf2, rfl⟩ ⟨g1, hg1, g2, hg2, rfl⟩ rw [PartialHomeomorph.prod_symm, PartialHomeomorph.prod_trans] - have h1 := (contDiffGroupoid ⊤ I).compatible hf1 hg1 - have h2 := (contDiffGroupoid ⊤ I').compatible hf2 hg2 + have h1 := (contDiffGroupoid ∞ I).compatible hf1 hg1 + have h2 := (contDiffGroupoid ∞ I').compatible hf2 hg2 exact contDiffGroupoid_prod h1 h2 end SmoothManifoldWithCorners @@ -1065,13 +1062,13 @@ open SmoothManifoldWithCorners theorem contDiffOn_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) : - ContDiffOn 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) ((f'.extend I).symm ≫ f.extend I).source := by + ContDiffOn 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) ((f'.extend I).symm ≫ f.extend I).source := by rw [extend_coord_change_source, I.image_eq] exact (StructureGroupoid.compatible_of_mem_maximalAtlas hf' hf).1 theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : E} (hx : x ∈ ((f'.extend I).symm ≫ f.extend I).source) : - ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) x := by + ContDiffWithinAt 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) (range I) x := by apply (contDiffOn_extend_coord_change hf hf' x hx).mono_of_mem_nhdsWithin rw [extend_coord_change_source] at hx ⊢ obtain ⟨z, hz, rfl⟩ := hx @@ -1079,7 +1076,7 @@ theorem contDiffWithinAt_extend_coord_change [ChartedSpace H M] (hf : f ∈ maxi theorem contDiffWithinAt_extend_coord_change' [ChartedSpace H M] (hf : f ∈ maximalAtlas I M) (hf' : f' ∈ maximalAtlas I M) {x : M} (hxf : x ∈ f.source) (hxf' : x ∈ f'.source) : - ContDiffWithinAt 𝕜 ⊤ (f.extend I ∘ (f'.extend I).symm) (range I) (f'.extend I x) := by + ContDiffWithinAt 𝕜 ∞ (f.extend I ∘ (f'.extend I).symm) (range I) (f'.extend I x) := by refine contDiffWithinAt_extend_coord_change hf hf' ?_ rw [← extend_image_source_inter] exact mem_image_of_mem _ ⟨hxf', hxf⟩ @@ -1408,15 +1405,14 @@ theorem ext_coord_change_source (x x' : M) : open SmoothManifoldWithCorners theorem contDiffOn_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) : - ContDiffOn 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) + ContDiffOn 𝕜 ∞ (extChartAt I x ∘ (extChartAt I x').symm) ((extChartAt I x').symm ≫ extChartAt I x).source := contDiffOn_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') theorem contDiffWithinAt_ext_coord_change [SmoothManifoldWithCorners I M] (x x' : M) {y : E} (hy : y ∈ ((extChartAt I x').symm ≫ extChartAt I x).source) : - ContDiffWithinAt 𝕜 ⊤ (extChartAt I x ∘ (extChartAt I x').symm) (range I) y := - contDiffWithinAt_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') - hy + ContDiffWithinAt 𝕜 ∞ (extChartAt I x ∘ (extChartAt I x').symm) (range I) y := + contDiffWithinAt_extend_coord_change (chart_mem_maximalAtlas x) (chart_mem_maximalAtlas x') hy variable (I I') in /-- Conjugating a function to write it in the preferred charts around `x`. diff --git a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean index 751bf15036342..60a27aee36286 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/SmoothSection.lean @@ -19,6 +19,9 @@ sections of a smooth vector bundle over a manifold `M` and prove that it's a mod open Bundle Filter Function open scoped Bundle Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] {H : Type*} [TopologicalSpace H] (I : ModelWithCorners 𝕜 E H) diff --git a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean index f40f1f950c07a..5931fefe4f64b 100644 --- a/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean +++ b/Mathlib/Geometry/Manifold/VectorBundle/Tangent.lean @@ -36,7 +36,7 @@ and a smooth manifold. open Bundle Set SmoothManifoldWithCorners PartialHomeomorph ContinuousLinearMap -open scoped Manifold Topology Bundle +open scoped Manifold Topology Bundle ContDiff noncomputable section @@ -49,7 +49,6 @@ variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] {E : Type*} [NormedAddCom [SmoothManifoldWithCorners I M] {M' : Type*} [TopologicalSpace M'] [ChartedSpace H' M'] [SmoothManifoldWithCorners I' M'] {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] - /-- Auxiliary lemma for tangent spaces: the derivative of a coordinate change between two charts is smooth on its source. -/ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : @@ -58,12 +57,12 @@ theorem contDiffOn_fderiv_coord_change (i j : atlas H M) : have h : ((i.1.extend I).symm ≫ j.1.extend I).source ⊆ range I := by rw [i.1.extend_coord_change_source]; apply image_subset_range intro x hx - refine (ContDiffWithinAt.fderivWithin_right ?_ I.uniqueDiffOn le_top <| h hx).mono h + refine (ContDiffWithinAt.fderivWithin_right ?_ I.uniqueDiffOn (n := ∞) (mod_cast le_top) + <| h hx).mono h refine (PartialHomeomorph.contDiffOn_extend_coord_change (subset_maximalAtlas j.2) (subset_maximalAtlas i.2) x hx).mono_of_mem_nhdsWithin ?_ exact i.1.extend_coord_change_source_mem_nhdsWithin j.1 hx - open SmoothManifoldWithCorners variable (I M) in @@ -104,9 +103,9 @@ def tangentBundleCore : VectorBundleCore 𝕜 M E (atlas H M) where simp_rw [Function.comp_apply, (j.1.extend I).left_inv hy] · simp_rw [Function.comp_apply, i.1.extend_left_inv hxi, j.1.extend_left_inv hxj] · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas k.2) - (subset_maximalAtlas j.2) hxk hxj).differentiableWithinAt le_top + (subset_maximalAtlas j.2) hxk hxj).differentiableWithinAt (mod_cast le_top) · exact (contDiffWithinAt_extend_coord_change' (subset_maximalAtlas j.2) - (subset_maximalAtlas i.2) hxj hxi).differentiableWithinAt le_top + (subset_maximalAtlas i.2) hxj hxi).differentiableWithinAt (mod_cast le_top) · intro x _; exact mem_range_self _ · exact I.uniqueDiffWithinAt_image · rw [Function.comp_apply, i.1.extend_left_inv hxi] diff --git a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean index 5551413ea3416..a6416ab861ef1 100644 --- a/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean +++ b/Mathlib/Geometry/Manifold/WhitneyEmbedding.lean @@ -33,6 +33,9 @@ variable {ι : Type uι} {E : Type uE} [NormedAddCommGroup E] [NormedSpace ℝ E open Function Filter Module Set Topology open scoped Manifold +/- Next line is necessary while the manifold smoothness class is not extended to `ω`. +Later, replace with `open scoped ContDiff`. -/ +local notation "∞" => (⊤ : ℕ∞) noncomputable section diff --git a/MathlibTest/fun_prop2.lean b/MathlibTest/fun_prop2.lean index e98b764c72c0a..8a5d906d3f74b 100644 --- a/MathlibTest/fun_prop2.lean +++ b/MathlibTest/fun_prop2.lean @@ -76,7 +76,7 @@ example : AEMeasurable T := by fun_prop -private theorem t1 : (5: ℕ) + (1 : ℕ∞) ≤ (12 : ℕ∞) := by norm_cast +private theorem t1 : (5: ℕ) + (1 : ℕ∞) ≤ (12 : WithTop ℕ∞) := by norm_cast example {f : ℝ → ℝ} (hf : ContDiff ℝ 12 f) : Differentiable ℝ (iteratedDeriv 5 (fun x => f (2*(f (x + x))) + x)) := by From beb3168d9638a2a180f3b33e673e3ac69444c965 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 25 Nov 2024 12:42:00 +0000 Subject: [PATCH 141/250] feat: additivize equivariance of morphisms of actions (#19171) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Additivize the definition of equivariant morphisms with respect to multiplicative actions. The notation introduced is `X →ₑᵥ[φ] Y` for equivariant morphisms, and `X →ᵥ[M] Y` when the underlying map is identity. The goal is to be able to fully additivize the properties of blocks for actions and the definition of primitive actions in #12052 Co-authored-by: leanprover-community-mathlib4-bot --- Mathlib/GroupTheory/GroupAction/Hom.lean | 101 ++++++++++++++++++----- 1 file changed, 80 insertions(+), 21 deletions(-) diff --git a/Mathlib/GroupTheory/GroupAction/Hom.lean b/Mathlib/GroupTheory/GroupAction/Hom.lean index 95bd4983abab0..2ce20e3e8812b 100644 --- a/Mathlib/GroupTheory/GroupAction/Hom.lean +++ b/Mathlib/GroupTheory/GroupAction/Hom.lean @@ -15,6 +15,7 @@ import Mathlib.Algebra.Group.Hom.CompTypeclasses * `MulActionHom φ X Y`, the type of equivariant functions from `X` to `Y`, where `φ : M → N` is a map, `M` acting on the type `X` and `N` acting on the type of `Y`. + `AddActionHom φ X Y` is its additive version. * `DistribMulActionHom φ A B`, the type of equivariant additive monoid homomorphisms from `A` to `B`, where `φ : M → N` is a morphism of monoids, @@ -25,7 +26,8 @@ import Mathlib.Algebra.Group.Hom.CompTypeclasses The above types have corresponding classes: * `MulActionHomClass F φ X Y` states that `F` is a type of bundled `X → Y` homs - which are `φ`-equivariant + which are `φ`-equivariant; + `AddActionHomClass F φ X Y` is its additive version. * `DistribMulActionHomClass F φ A B` states that `F` is a type of bundled `A → B` homs preserving the additive monoid structure and `φ`-equivariant * `SMulSemiringHomClass F φ R S` states that `F` is a type of bundled `R → S` homs @@ -35,14 +37,19 @@ The above types have corresponding classes: We introduce the following notation to code equivariant maps (the subscript index `ₑ` is for *equivariant*) : -* `X →ₑ[φ] Y` is `MulActionHom φ X Y`. +* `X →ₑ[φ] Y` is `MulActionHom φ X Y` and `AddActionHom φ X Y` * `A →ₑ+[φ] B` is `DistribMulActionHom φ A B`. * `R →ₑ+*[φ] S` is `MulSemiringActionHom φ R S`. When `M = N` and `φ = MonoidHom.id M`, we provide the backward compatible notation : -* `X →[M] Y` is `MulActionHom (@id M) X Y` +* `X →[M] Y` is `MulActionHom (@id M) X Y` and `AddActionHom (@id M) X Y` * `A →+[M] B` is `DistribMulActionHom (MonoidHom.id M) A B` * `R →+*[M] S` is `MulSemiringActionHom (MonoidHom.id M) R S` + +The notation for `MulActionHom` and `AddActionHom` is the same, because it is unlikely +that it could lead to confusion — unless one needs types `M` and `X` with simultaneous +instances of `Mul M`, `Add M`, `SMul M X` and `VAdd M X`… + -/ assert_not_exists Submonoid @@ -56,11 +63,21 @@ variable (X : Type*) [SMul M X] [SMul M' X] variable (Y : Type*) [SMul N Y] [SMul M' Y] variable (Z : Type*) [SMul P Z] +/-- Equivariant functions : +When `φ : M → N` is a function, and types `X` and `Y` are endowed with additive actions +of `M` and `N`, a function `f : X → Y` is `φ`-equivariant if `f (m +ᵥ x) = (φ m) +ᵥ (f x)`. -/ +structure AddActionHom {M N : Type*} (φ: M → N) (X : Type*) [VAdd M X] (Y : Type*) [VAdd N Y] where + /-- The underlying function. -/ + protected toFun : X → Y + /-- The proposition that the function commutes with the additive actions. -/ + protected map_vadd' : ∀ (m : M) (x : X), toFun (m +ᵥ x) = (φ m) +ᵥ toFun x + /-- Equivariant functions : When `φ : M → N` is a function, and types `X` and `Y` are endowed with actions of `M` and `N`, a function `f : X → Y` is `φ`-equivariant if `f (m • x) = (φ m) • (f x)`. -/ -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): this linter isn't ported yet. -- @[nolint has_nonempty_instance] +@[to_additive] structure MulActionHom where /-- The underlying function. -/ protected toFun : X → Y @@ -70,20 +87,40 @@ structure MulActionHom where /- Porting note: local notation given a name, conflict with Algebra.Hom.GroupAction see https://github.com/leanprover/lean4/issues/2000 -/ /-- `φ`-equivariant functions `X → Y`, -where `φ : M → N`, where `M` and `N` act on `X` and `Y` respectively -/ +where `φ : M → N`, where `M` and `N` act on `X` and `Y` respectively.-/ notation:25 (name := «MulActionHomLocal≺») X " →ₑ[" φ:25 "] " Y:0 => MulActionHom φ X Y -/-- `M`-equivariant functions `X → Y` with respect to the action of `M` - -This is the same as `X →ₑ[@id M] Y` -/ +/-- `M`-equivariant functions `X → Y` with respect to the action of `M`. +This is the same as `X →ₑ[@id M] Y`. -/ notation:25 (name := «MulActionHomIdLocal≺») X " →[" M:25 "] " Y:0 => MulActionHom (@id M) X Y +/-- `φ`-equivariant functions `X → Y`, +where `φ : M → N`, where `M` and `N` act additively on `X` and `Y` respectively +We use the same notation as for multiplicative actions, as conflicts are unlikely. -/ +notation:25 (name := «AddActionHomLocal≺») X " →ₑ[" φ:25 "] " Y:0 => AddActionHom φ X Y + +/-- `M`-equivariant functions `X → Y` with respect to the additive action of `M`. +This is the same as `X →ₑ[@id M] Y`. + +We use the same notation as for multiplicative actions, as conflicts are unlikely. -/ +notation:25 (name := «AddActionHomIdLocal≺») X " →[" M:25 "] " Y:0 => AddActionHom (@id M) X Y + +/-- `AddActionSemiHomClass F φ X Y` states that + `F` is a type of morphisms which are `φ`-equivariant. + +You should extend this class when you extend `AddActionHom`. -/ +class AddActionSemiHomClass (F : Type*) + {M N : outParam Type*} (φ : outParam (M → N)) + (X Y : outParam Type*) [VAdd M X] [VAdd N Y] [FunLike F X Y] : Prop where + /-- The proposition that the function preserves the action. -/ + map_vaddₛₗ : ∀ (f : F) (c : M) (x : X), f (c +ᵥ x) = (φ c) +ᵥ (f x) /-- `MulActionSemiHomClass F φ X Y` states that `F` is a type of morphisms which are `φ`-equivariant. You should extend this class when you extend `MulActionHom`. -/ +@[to_additive] class MulActionSemiHomClass (F : Type*) {M N : outParam Type*} (φ : outParam (M → N)) (X Y : outParam Type*) [SMul M X] [SMul N Y] [FunLike F X Y] : Prop where @@ -91,19 +128,23 @@ class MulActionSemiHomClass (F : Type*) map_smulₛₗ : ∀ (f : F) (c : M) (x : X), f (c • x) = (φ c) • (f x) export MulActionSemiHomClass (map_smulₛₗ) +export AddActionSemiHomClass (map_vaddₛₗ) /-- `MulActionHomClass F M X Y` states that `F` is a type of morphisms which are equivariant with respect to actions of `M` This is an abbreviation of `MulActionSemiHomClass`. -/ +@[to_additive "`MulActionHomClass F M X Y` states that `F` is a type of +morphisms which are equivariant with respect to actions of `M` +This is an abbreviation of `MulActionSemiHomClass`."] abbrev MulActionHomClass (F : Type*) (M : outParam Type*) (X Y : outParam Type*) [SMul M X] [SMul M Y] [FunLike F X Y] := MulActionSemiHomClass F (@id M) X Y -instance : FunLike (MulActionHom φ X Y) X Y where +@[to_additive] instance : FunLike (MulActionHom φ X Y) X Y where coe := MulActionHom.toFun coe_injective' f g h := by cases f; cases g; congr -@[simp] +@[to_additive (attr := simp)] theorem map_smul {F M X Y : Type*} [SMul M X] [SMul M Y] [FunLike F X Y] [MulActionHomClass F M X Y] (f : F) (c : M) (x : X) : f (c • x) = c • f x := @@ -113,10 +154,12 @@ theorem map_smul {F M X Y : Type*} [SMul M X] [SMul M Y] -- Porting note: removed has_coe_to_fun instance, coercions handled differently now +@[to_additive] instance : MulActionSemiHomClass (X →ₑ[φ] Y) φ X Y where map_smulₛₗ := MulActionHom.map_smul' initialize_simps_projections MulActionHom (toFun → apply) +initialize_simps_projections AddActionHom (toFun → apply) namespace MulActionHom @@ -128,7 +171,10 @@ see also Algebra.Hom.Group -/ /-- Turn an element of a type `F` satisfying `MulActionSemiHomClass F φ X Y` into an actual `MulActionHom`. This is declared as the default coercion from `F` to `MulActionSemiHom φ X Y`. -/ -@[coe] +@[to_additive (attr := coe) + "Turn an element of a type `F` satisfying `AddActionSemiHomClass F φ X Y` + into an actual `AddActionHom`. + This is declared as the default coercion from `F` to `AddActionSemiHom φ X Y`."] def _root_.MulActionSemiHomClass.toMulActionHom [MulActionSemiHomClass F φ X Y] (f : F) : X →ₑ[φ] Y where toFun := DFunLike.coe f @@ -136,39 +182,44 @@ def _root_.MulActionSemiHomClass.toMulActionHom [MulActionSemiHomClass F φ X Y] /-- Any type satisfying `MulActionSemiHomClass` can be cast into `MulActionHom` via `MulActionHomSemiClass.toMulActionHom`. -/ +@[to_additive] instance [MulActionSemiHomClass F φ X Y] : CoeTC F (X →ₑ[φ] Y) := ⟨MulActionSemiHomClass.toMulActionHom⟩ variable (M' X Y F) in /-- If Y/X/M forms a scalar tower, any map X → Y preserving X-action also preserves M-action. -/ +@[to_additive] theorem _root_.IsScalarTower.smulHomClass [MulOneClass X] [SMul X Y] [IsScalarTower M' X Y] [MulActionHomClass F X X Y] : MulActionHomClass F M' X Y where map_smulₛₗ f m x := by rw [← mul_one (m • x), ← smul_eq_mul, map_smul, smul_assoc, ← map_smul, smul_eq_mul, mul_one, id_eq] +@[to_additive] protected theorem map_smul (f : X →[M'] Y) (m : M') (x : X) : f (m • x) = m • f x := map_smul f m x -@[ext] +@[to_additive (attr := ext)] theorem ext {f g : X →ₑ[φ] Y} : (∀ x, f x = g x) → f = g := DFunLike.ext f g +@[to_additive] protected theorem congr_fun {f g : X →ₑ[φ] Y} (h : f = g) (x : X) : f x = g x := DFunLike.congr_fun h _ /-- Two equal maps on scalars give rise to an equivariant map for identity -/ +@[to_additive "Two equal maps on scalars give rise to an equivariant map for identity"] def ofEq {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) : X →ₑ[φ'] Y where toFun := f.toFun map_smul' m a := h ▸ f.map_smul' m a -@[simp] +@[to_additive (attr := simp)] theorem ofEq_coe {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) : (f.ofEq h).toFun = f.toFun := rfl -@[simp] +@[to_additive (attr := simp)] theorem ofEq_apply {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) (a : X) : (f.ofEq h) a = f a := rfl @@ -177,12 +228,13 @@ theorem ofEq_apply {φ' : M → N} (h : φ = φ') (f : X →ₑ[φ] Y) (a : X) : variable {ψ χ} (M N) /-- The identity map as an equivariant map. -/ +@[to_additive "The identity map as an equivariant map."] protected def id : X →[M] X := ⟨id, fun _ _ => rfl⟩ variable {M N Z} -@[simp] +@[to_additive (attr := simp)] theorem id_apply (x : X) : MulActionHom.id M x = x := rfl @@ -197,6 +249,7 @@ variable {φ ψ χ X Y Z} -- attribute [instance] CompTriple.id_comp CompTriple.comp_id /-- Composition of two equivariant maps. -/ +@[to_additive "Composition of two equivariant maps."] def comp (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [κ : CompTriple φ ψ χ] : X →ₑ[χ] Z := ⟨g ∘ f, fun m x => @@ -206,22 +259,22 @@ def comp (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [κ : CompTriple φ ψ χ] : _ = (ψ ∘ φ) m • g (f x) := rfl _ = χ m • g (f x) := by rw [κ.comp_eq] ⟩ -@[simp] +@[to_additive (attr := simp)] theorem comp_apply (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) [CompTriple φ ψ χ] (x : X) : g.comp f x = g (f x) := rfl -@[simp] +@[to_additive (attr := simp)] theorem id_comp (f : X →ₑ[φ] Y) : (MulActionHom.id N).comp f = f := ext fun x => by rw [comp_apply, id_apply] -@[simp] +@[to_additive (attr := simp)] theorem comp_id (f : X →ₑ[φ] Y) : f.comp (MulActionHom.id M) = f := ext fun x => by rw [comp_apply, id_apply] -@[simp] +@[to_additive (attr := simp)] theorem comp_assoc {Q T : Type*} [SMul Q T] {η : P → Q} {θ : M → Q} {ζ : N → Q} (h : Z →ₑ[η] T) (g : Y →ₑ[ψ] Z) (f : X →ₑ[φ] Y) @@ -232,8 +285,9 @@ theorem comp_assoc {Q T : Type*} [SMul Q T] variable {φ' : N → M} variable {Y₁ : Type*} [SMul M Y₁] + /-- The inverse of a bijective equivariant map is equivariant. -/ -@[simps] +@[to_additive (attr := simps) "The inverse of a bijective equivariant map is equivariant."] def inverse (f : X →[M] Y₁) (g : Y₁ → X) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : Y₁ →[M] X where toFun := g @@ -245,7 +299,7 @@ def inverse (f : X →[M] Y₁) (g : Y₁ → X) /-- The inverse of a bijective equivariant map is equivariant. -/ -@[simps] +@[to_additive (attr := simps) "The inverse of a bijective equivariant map is equivariant."] def inverse' (f : X →ₑ[φ] Y) (g : Y → X) (k : Function.RightInverse φ' φ) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : Y →ₑ[φ'] X where @@ -257,11 +311,13 @@ def inverse' (f : X →ₑ[φ] Y) (g : Y → X) (k : Function.RightInverse φ' _ = g (f (φ' m • g x)) := by rw [map_smulₛₗ] _ = φ' m • g x := by rw [h₁] +@[to_additive] lemma inverse_eq_inverse' (f : X →[M] Y₁) (g : Y₁ → X) (h₁ : Function.LeftInverse g f) (h₂ : Function.RightInverse g f) : inverse f g h₁ h₂ = inverse' f g (congrFun rfl) h₁ h₂ := by rfl +@[to_additive] theorem inverse'_inverse' {f : X →ₑ[φ] Y} {g : Y → X} {k₁ : Function.LeftInverse φ' φ} {k₂ : Function.RightInverse φ' φ} @@ -269,6 +325,7 @@ theorem inverse'_inverse' inverse' (inverse' f g k₂ h₁ h₂) f k₁ h₂ h₁ = f := ext fun _ => rfl +@[to_additive] theorem comp_inverse' {f : X →ₑ[φ] Y} {g : Y → X} {k₁ : Function.LeftInverse φ' φ} {k₂ : Function.RightInverse φ' φ} {h₁ : Function.LeftInverse g f} {h₂ : Function.RightInverse g f} : @@ -279,6 +336,7 @@ theorem comp_inverse' {f : X →ₑ[φ] Y} {g : Y → X} simp only [comp_apply, inverse_apply, id_apply] exact h₁ x +@[to_additive] theorem inverse'_comp {f : X →ₑ[φ] Y} {g : Y → X} {k₂ : Function.RightInverse φ' φ} {h₁ : Function.LeftInverse g f} {h₂ : Function.RightInverse g f} : @@ -290,7 +348,8 @@ theorem inverse'_comp {f : X →ₑ[φ] Y} {g : Y → X} /-- If actions of `M` and `N` on `α` commute, then for `c : M`, `(c • · : α → α)` is an `N`-action homomorphism. -/ -@[simps] +@[to_additive (attr := simps) "If additive actions of `M` and `N` on `α` commute, + then for `c : M`, `(c • · : α → α)` is an `N`-additive action homomorphism."] def _root_.SMulCommClass.toMulActionHom {M} (N α : Type*) [SMul M α] [SMul N α] [SMulCommClass M N α] (c : M) : α →[N] α where From 7f28d7f73f74b31f72449bb9eb240db3ca68fe38 Mon Sep 17 00:00:00 2001 From: Antoine Chambert-Loir Date: Mon, 25 Nov 2024 12:42:01 +0000 Subject: [PATCH 142/250] feat(Data/Finsupp/MonomialOrder): monomial orders (#19453) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Definition of monomial orders, with the example of the lexicographic ordering This is the starting point of implementation of the beginnings of Gröbner theory. 3 subsequent PRs will address: - division algorithm - example of the homogeneous lexicographic ordering - example of the homogeneous reverse lexicographic ordering These various orderings (and other ones…) may look exotic but are actually used in algebraic geometry. --- Mathlib.lean | 1 + Mathlib/Data/Finsupp/MonomialOrder.lean | 156 ++++++++++++++++++++++++ docs/references.bib | 15 +++ 3 files changed, 172 insertions(+) create mode 100644 Mathlib/Data/Finsupp/MonomialOrder.lean diff --git a/Mathlib.lean b/Mathlib.lean index 34090017c1344..f73b287687527 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2449,6 +2449,7 @@ import Mathlib.Data.Finsupp.Fintype import Mathlib.Data.Finsupp.Indicator import Mathlib.Data.Finsupp.Interval import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.MonomialOrder import Mathlib.Data.Finsupp.Multiset import Mathlib.Data.Finsupp.NeLocus import Mathlib.Data.Finsupp.Notation diff --git a/Mathlib/Data/Finsupp/MonomialOrder.lean b/Mathlib/Data/Finsupp/MonomialOrder.lean new file mode 100644 index 0000000000000..431fcb7d3ce34 --- /dev/null +++ b/Mathlib/Data/Finsupp/MonomialOrder.lean @@ -0,0 +1,156 @@ +/- +Copyright (c) 2024 Antoine Chambert-Loir. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Antoine Chambert-Loir +-/ +import Mathlib.Data.Finsupp.Lex +import Mathlib.Data.Finsupp.WellFounded +import Mathlib.Data.List.TFAE + +/-! # Monomial orders + +## Monomial orders + +A *monomial order* is well ordering relation on a type of the form `σ →₀ ℕ` which +is compatible with addition and for which `0` is the smallest element. +Since several monomial orders may have to be used simultaneously, one cannot +get them as instances. +In this formalization, they are presented as a structure `MonomialOrder` which encapsulates +`MonomialOrder.toSyn`, an additive and monotone isomorphism to a linearly ordered cancellative +additive commutative monoid. +The entry `MonomialOrder.wf` asserts that `MonomialOrder.syn` is well founded. + +The terminology comes from commutative algebra and algebraic geometry, especially Gröbner bases, +where `c : σ →₀ ℕ` are exponents of monomials. + +Given a monomial order `m : MonomialOrder σ`, we provide the notation +`c ≼[m] d` and `c ≺[m] d` to compare `c d : σ →₀ ℕ` with respect to `m`. +It is activated using `open scoped MonomialOrder`. + +## Examples + +Commutative algebra defines many monomial orders, with different usefulness ranges. +In this file, we provide the basic example of lexicographic ordering. +For the graded lexicographic ordering, see `Mathlib/Data/Finsupp/DegLex.lean` + +* `MonomialOrder.lex` : the lexicographic ordering on `σ →₀ ℕ`. +For this, `σ` needs to be embedded with an ordering relation which satisfies `WellFoundedGT σ`. +(This last property is automatic when `σ` is finite). + +The type synonym is `Lex (σ →₀ ℕ)` and the two lemmas `MonomialOrder.lex_le_iff` +and `MonomialOrder.lex_lt_iff` rewrite the ordering as comparisons in the type `Lex (σ →₀ ℕ)`. + +## References + +* [Cox, Little and O'Shea, *Ideals, varieties, and algorithms*][coxlittleoshea1997] +* [Becker and Weispfenning, *Gröbner bases*][Becker-Weispfenning1993] + +## Note + +In algebraic geometry, when the finitely many variables are indexed by integers, +it is customary to order them using the opposite order : `MvPolynomial.X 0 > MvPolynomial.X 1 > … ` + +-/ + +/-- Monomial orders : equivalence of `σ →₀ ℕ` with a well ordered type -/ +structure MonomialOrder (σ : Type*) where + /-- The synonym type -/ + syn : Type* + /-- `syn` is a linearly ordered cancellative additive commutative monoid -/ + locacm : LinearOrderedCancelAddCommMonoid syn := by infer_instance + /-- the additive equivalence from `σ →₀ ℕ` to `syn` -/ + toSyn : (σ →₀ ℕ) ≃+ syn + /-- `toSyn` is monotone -/ + toSyn_monotone : Monotone toSyn + /-- `syn` is a well ordering -/ + wf : WellFoundedLT syn := by infer_instance + +attribute [instance] MonomialOrder.locacm MonomialOrder.wf + +namespace MonomialOrder + +variable {σ : Type*} (m : MonomialOrder σ) + +lemma le_add_right (a b : σ →₀ ℕ) : + m.toSyn a ≤ m.toSyn a + m.toSyn b := by + rw [← map_add] + exact m.toSyn_monotone le_self_add + +instance orderBot : OrderBot (m.syn) where + bot := 0 + bot_le a := by + have := m.le_add_right 0 (m.toSyn.symm a) + simp [map_add, zero_add] at this + exact this + +@[simp] +theorem bot_eq_zero : (⊥ : m.syn) = 0 := rfl + +theorem eq_zero_iff {a : m.syn} : a = 0 ↔ a ≤ 0 := eq_bot_iff + +lemma toSyn_strictMono : StrictMono (m.toSyn) := by + apply m.toSyn_monotone.strictMono_of_injective m.toSyn.injective + +/-- Given a monomial order, notation for the corresponding strict order relation on `σ →₀ ℕ` -/ +scoped +notation:25 c "≺[" m:25 "]" d:25 => (MonomialOrder.toSyn m c < MonomialOrder.toSyn m d) + +/-- Given a monomial order, notation for the corresponding order relation on `σ →₀ ℕ` -/ +scoped +notation:25 c "≼[" m:25 "]" d:25 => (MonomialOrder.toSyn m c ≤ MonomialOrder.toSyn m d) + +end MonomialOrder + +section Lex + +open Finsupp + +open scoped MonomialOrder + +-- The linear order on `Finsupp`s obtained by the lexicographic ordering. -/ +noncomputable instance {α N : Type*} [LinearOrder α] [OrderedCancelAddCommMonoid N] : + OrderedCancelAddCommMonoid (Lex (α →₀ N)) where + le_of_add_le_add_left a b c h := by simpa only [add_le_add_iff_left] using h + add_le_add_left a b h c := by simpa only [add_le_add_iff_left] using h + +theorem Finsupp.lex_lt_iff {α N : Type*} [LinearOrder α] [LinearOrder N] [Zero N] + {a b : Lex (α →₀ N)} : + a < b ↔ ∃ i, (∀ j, j< i → ofLex a j = ofLex b j) ∧ ofLex a i < ofLex b i := + Finsupp.lex_def + +theorem Finsupp.lex_le_iff {α N : Type*} [LinearOrder α] [LinearOrder N] [Zero N] + {a b : Lex (α →₀ N)} : + a ≤ b ↔ a = b ∨ ∃ i, (∀ j, j< i → ofLex a j = ofLex b j) ∧ ofLex a i < ofLex b i := by + rw [le_iff_eq_or_lt, Finsupp.lex_lt_iff] + +/-- for the lexicographic ordering, X 0 * X 1 < X 0 ^ 2 -/ +example : toLex (Finsupp.single 0 2) > toLex (Finsupp.single 0 1 + Finsupp.single 1 1) := by + use 0; simp + +/-- for the lexicographic ordering, X 1 < X 0 -/ +example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 1) := by + use 0; simp + +/-- for the lexicographic ordering, X 1 < X 0 ^ 2 -/ +example : toLex (Finsupp.single 1 1) < toLex (Finsupp.single 0 2) := by + use 0; simp + + +variable {σ : Type*} [LinearOrder σ] + +/-- The lexicographic order on `σ →₀ ℕ`, as a `MonomialOrder` -/ +noncomputable def MonomialOrder.lex [WellFoundedGT σ] : + MonomialOrder σ where + syn := Lex (σ →₀ ℕ) + toSyn := + { toEquiv := toLex + map_add' := toLex_add } + toSyn_monotone := Finsupp.toLex_monotone + +theorem MonomialOrder.lex_le_iff [WellFoundedGT σ] {c d : σ →₀ ℕ} : + c ≼[lex] d ↔ toLex c ≤ toLex d := Iff.rfl + +theorem MonomialOrder.lex_lt_iff [WellFoundedGT σ] {c d : σ →₀ ℕ} : + c ≺[lex] d ↔ toLex c < toLex d := Iff.rfl + +end Lex diff --git a/docs/references.bib b/docs/references.bib index 3d890cec41f2a..0e5ce6b9bd094 100644 --- a/docs/references.bib +++ b/docs/references.bib @@ -247,6 +247,21 @@ @Book{ beals2004 year = {2004} } +@Book{ Becker-Weispfenning1993, + address = {New York, NY}, + series = {Graduate Texts in Mathematics}, + title = {Gröbner Bases}, + volume = {141}, + isbn = {978-1-4612-6944-1}, + url = {http://link.springer.com/10.1007/978-1-4612-0913-3}, + doi = {10.1007/978-1-4612-0913-3}, + publisher = {Springer New York}, + author = {Becker, Thomas and Weispfenning, Volker}, + year = {1993}, + collection = {Graduate Texts in Mathematics}, + language = {en} +} + @Book{ behrends1979, author = {Ehrhard {Behrends}}, title = {{M-structure and the Banach-Stone theorem}}, From f26c0a21e7f81a4e95e0703e492af2d729da839d Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Mon, 25 Nov 2024 13:12:47 +0000 Subject: [PATCH 143/250] chore(RingTheory/Noetherian): further split `Defs.lean` (#19364) Now that we have split `RingTheory.Finiteness`, we have the ability to clean up the import hierarchy of `RingTheory.PrincipalIdealDomain` further by splitting `Noetherian/Defs.lean` further. Now `Noetherian/Defs.lean` only contains results relating finite generation of submodule and wellfoundedness of the inclusion relation. --- Mathlib.lean | 1 + .../Algebra/Module/FinitePresentation.lean | 2 +- Mathlib/FieldTheory/Tower.lean | 2 +- .../RingTheory/GradedAlgebra/Noetherian.lean | 2 +- .../RingTheory/Ideal/Quotient/Noetherian.lean | 2 +- Mathlib/RingTheory/Noetherian/Basic.lean | 366 ++++++++++++++++++ Mathlib/RingTheory/Noetherian/Defs.lean | 298 +------------- Mathlib/RingTheory/Noetherian/Orzech.lean | 2 +- Mathlib/RingTheory/Polynomial/Basic.lean | 2 +- Mathlib/RingTheory/PrimeSpectrum.lean | 2 +- Mathlib/RingTheory/PrincipalIdealDomain.lean | 1 + 11 files changed, 379 insertions(+), 301 deletions(-) create mode 100644 Mathlib/RingTheory/Noetherian/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index f73b287687527..39dccdce392ff 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -4385,6 +4385,7 @@ import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.Nilpotent.Basic import Mathlib.RingTheory.Nilpotent.Defs import Mathlib.RingTheory.Nilpotent.Lemmas +import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.Noetherian.Defs import Mathlib.RingTheory.Noetherian.Filter import Mathlib.RingTheory.Noetherian.Nilpotent diff --git a/Mathlib/Algebra/Module/FinitePresentation.lean b/Mathlib/Algebra/Module/FinitePresentation.lean index 553e1aee1ced2..b5737ac1b8e49 100644 --- a/Mathlib/Algebra/Module/FinitePresentation.lean +++ b/Mathlib/Algebra/Module/FinitePresentation.lean @@ -9,7 +9,7 @@ import Mathlib.LinearAlgebra.Isomorphisms import Mathlib.RingTheory.Finiteness.Projective import Mathlib.RingTheory.Finiteness.TensorProduct import Mathlib.RingTheory.Localization.BaseChange -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! diff --git a/Mathlib/FieldTheory/Tower.lean b/Mathlib/FieldTheory/Tower.lean index dbe2c6efe3f02..59cdb775f2ec4 100644 --- a/Mathlib/FieldTheory/Tower.lean +++ b/Mathlib/FieldTheory/Tower.lean @@ -3,7 +3,7 @@ Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Finiteness of `IsScalarTower` diff --git a/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean b/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean index cba8dab196628..9bd8cc5122295 100644 --- a/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean +++ b/Mathlib/RingTheory/GradedAlgebra/Noetherian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Fangming Li -/ import Mathlib.RingTheory.GradedAlgebra.Basic -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # The properties of a graded Noetherian ring. diff --git a/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean b/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean index abe698078c199..2f2f268de1998 100644 --- a/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean +++ b/Mathlib/RingTheory/Ideal/Quotient/Noetherian.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.RingTheory.Ideal.Quotient.Operations -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Noetherian quotient rings and quotient modules diff --git a/Mathlib/RingTheory/Noetherian/Basic.lean b/Mathlib/RingTheory/Noetherian/Basic.lean new file mode 100644 index 0000000000000..38156a143989d --- /dev/null +++ b/Mathlib/RingTheory/Noetherian/Basic.lean @@ -0,0 +1,366 @@ +/- +Copyright (c) 2018 Mario Carneiro, Kevin Buzzard. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Kevin Buzzard +-/ +import Mathlib.LinearAlgebra.Quotient.Basic +import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Finiteness.Cardinality +import Mathlib.RingTheory.Finiteness.Finsupp +import Mathlib.RingTheory.Ideal.Maps + +/-! +# Noetherian rings and modules + +The following are equivalent for a module M over a ring R: +1. Every increasing chain of submodules M₁ ⊆ M₂ ⊆ M₃ ⊆ ⋯ eventually stabilises. +2. Every submodule is finitely generated. + +A module satisfying these equivalent conditions is said to be a *Noetherian* R-module. +A ring is a *Noetherian ring* if it is Noetherian as a module over itself. + +(Note that we do not assume yet that our rings are commutative, +so perhaps this should be called "left Noetherian". +To avoid cumbersome names once we specialize to the commutative case, +we don't make this explicit in the declaration names.) + +## Main definitions + +Let `R` be a ring and let `M` and `P` be `R`-modules. Let `N` be an `R`-submodule of `M`. + +* `IsNoetherian R M` is the proposition that `M` is a Noetherian `R`-module. It is a class, + implemented as the predicate that all `R`-submodules of `M` are finitely generated. + +## Main statements + +* `isNoetherian_iff` is the theorem that an R-module M is Noetherian iff `>` is well-founded on + `Submodule R M`. + +Note that the Hilbert basis theorem, that if a commutative ring R is Noetherian then so is R[X], +is proved in `RingTheory.Polynomial`. + +## References + +* [M. F. Atiyah and I. G. Macdonald, *Introduction to commutative algebra*][atiyah-macdonald] +* [samuel1967] + +## Tags + +Noetherian, noetherian, Noetherian ring, Noetherian module, noetherian ring, noetherian module + +-/ + +assert_not_exists Matrix + +open Set Pointwise + +section + +variable {R : Type*} {M : Type*} {P : Type*} +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid P] +variable [Module R M] [Module R P] + +open IsNoetherian + +variable (M) + +theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = ⊤) [IsNoetherian R M] : + IsNoetherian R P := + ⟨fun s => + have : (s.comap f).map f = s := Submodule.map_comap_eq_self <| hf.symm ▸ le_top + this ▸ (noetherian _).map _⟩ + +variable {M} + +instance isNoetherian_range (f : M →ₗ[R] P) [IsNoetherian R M] : + IsNoetherian R (LinearMap.range f) := + isNoetherian_of_surjective _ _ f.range_rangeRestrict + +instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] + [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : + IsNoetherian R (M ⧸ N) := + isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| + LinearMap.range_eq_top.mpr N.mkQ_surjective + +@[deprecated (since := "2024-04-27"), nolint defLemma] +alias Submodule.Quotient.isNoetherian := isNoetherian_quotient + +theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := + isNoetherian_of_surjective _ f.toLinearMap f.range + +theorem LinearEquiv.isNoetherian_iff (f : M ≃ₗ[R] P) : IsNoetherian R M ↔ IsNoetherian R P := + ⟨fun _ ↦ isNoetherian_of_linearEquiv f, fun _ ↦ isNoetherian_of_linearEquiv f.symm⟩ + +theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := + Submodule.topEquiv.isNoetherian_iff + +theorem isNoetherian_of_injective [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) : + IsNoetherian R M := + isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f hf).symm + +theorem fg_of_injective [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) + (hf : Function.Injective f) : N.FG := + haveI := isNoetherian_of_injective f hf + IsNoetherian.noetherian N + +end + +namespace Module + +variable {R M N : Type*} +variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] +variable (R M) + +-- see Note [lower instance priority] +instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := + ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ + +-- see Note [lower instance priority] +instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := + ⟨IsNoetherian.noetherian ⊤⟩ + +instance {R₁ S : Type*} [CommSemiring R₁] [Semiring S] [Algebra R₁ S] + [IsNoetherian R₁ S] (I : Ideal S) : Module.Finite R₁ I := + IsNoetherian.finite R₁ ((I : Submodule S S).restrictScalars R₁) + +variable {R M} + +theorem Finite.of_injective [IsNoetherian R N] (f : M →ₗ[R] N) (hf : Function.Injective f) : + Module.Finite R M := + ⟨fg_of_injective f hf⟩ + +end Module + +section + +variable {R : Type*} {M : Type*} {P : Type*} +variable [Ring R] [AddCommGroup M] [AddCommGroup P] +variable [Module R M] [Module R P] + +open IsNoetherian + +theorem isNoetherian_of_ker_bot [IsNoetherian R P] (f : M →ₗ[R] P) (hf : LinearMap.ker f = ⊥) : + IsNoetherian R M := + isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f <| LinearMap.ker_eq_bot.mp hf).symm + +theorem fg_of_ker_bot [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) + (hf : LinearMap.ker f = ⊥) : N.FG := + haveI := isNoetherian_of_ker_bot f hf + IsNoetherian.noetherian N + +instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian R (M × P) := + ⟨fun s => + Submodule.fg_of_fg_map_of_fg_inf_ker (LinearMap.snd R M P) (noetherian _) <| + have : s ⊓ LinearMap.ker (LinearMap.snd R M P) ≤ LinearMap.range (LinearMap.inl R M P) := + fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ + Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ + +instance isNoetherian_sup (M₁ M₂ : Submodule R P) [IsNoetherian R M₁] [IsNoetherian R M₂] : + IsNoetherian R ↥(M₁ ⊔ M₂) := by + have := isNoetherian_range (M₁.subtype.coprod M₂.subtype) + rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this + +variable {ι : Type*} [Finite ι] + +instance isNoetherian_pi : + ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] + [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) + · infer_instance + · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm + +/-- A version of `isNoetherian_pi` for non-dependent functions. We need this instance because +sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to +prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ +instance isNoetherian_pi' [IsNoetherian R M] : IsNoetherian R (ι → M) := + isNoetherian_pi + +instance isNoetherian_iSup : + ∀ {M : ι → Submodule R P} [∀ i, IsNoetherian R (M i)], IsNoetherian R ↥(⨆ i, M i) := by + apply Finite.induction_empty_option _ _ _ ι + · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h + · intros; rw [iSup_of_empty]; infer_instance + · intro _ _ ih _ _; rw [iSup_option]; infer_instance + +end + +section CommRing + +variable (R M N : Type*) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] + [IsNoetherian R M] [Module.Finite R N] + +instance isNoetherian_linearMap_pi {ι : Type*} [Finite ι] : IsNoetherian R ((ι → R) →ₗ[R] M) := + let _i : Fintype ι := Fintype.ofFinite ι; isNoetherian_of_linearEquiv (Module.piEquiv ι R M) + +instance isNoetherian_linearMap : IsNoetherian R (N →ₗ[R] M) := by + obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R N + let g : (N →ₗ[R] M) →ₗ[R] (Fin n → R) →ₗ[R] M := (LinearMap.llcomp R (Fin n → R) N M).flip f + exact isNoetherian_of_injective g hf.injective_linearMapComp_right + +end CommRing + +open IsNoetherian Submodule Function + +section + +universe w + +variable {R M P : Type*} {N : Type w} [Semiring R] [AddCommMonoid M] [Module R M] [AddCommMonoid N] + [Module R N] [AddCommMonoid P] [Module R P] + +/-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ +theorem IsNoetherian.induction [IsNoetherian R M] {P : Submodule R M → Prop} + (hgt : ∀ I, (∀ J > I, P J) → P I) (I : Submodule R M) : P I := + IsWellFounded.induction _ I hgt + +end + +section + +universe w + +variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] + [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] + +lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} + (h : iSupIndep N) : + Set.Finite {i | N i ≠ ⊥} := + WellFoundedGT.finite_ne_bot_of_iSupIndep h + +@[deprecated (since := "2024-11-24")] +alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep + +/-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if +the module is Noetherian. -/ +theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} + (hv : LinearIndependent R v) : Finite ι := by + refine WellFoundedGT.finite_of_iSupIndep + hv.iSupIndep_span_singleton + fun i contra => ?_ + apply hv.ne_zero i + have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) + rwa [contra, Submodule.mem_bot] at this + +theorem LinearIndependent.set_finite_of_isNoetherian [Nontrivial R] {s : Set M} + (hi : LinearIndependent R ((↑) : s → M)) : s.Finite := + @Set.toFinite _ _ hi.finite_of_isNoetherian + +/-- If the first and final modules in an exact sequence are Noetherian, + then the middle module is also Noetherian. -/ +theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] + (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : + IsNoetherian R N := + isNoetherian_mk <| + wellFounded_gt_exact_sequence + (LinearMap.range f) + (Submodule.map (f.ker.liftQ f le_rfl)) + (Submodule.comap (f.ker.liftQ f le_rfl)) + (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) + (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) + (Submodule.giMapComap g.surjective_rangeRestrict) + (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) + (by simp [Submodule.comap_map_eq, h]) + +theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : + IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by + refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ + apply isNoetherian_of_range_eq_ker S.subtype S.mkQ + rw [Submodule.ker_mkQ, Submodule.range_subtype] + +/-- A sequence `f` of submodules of a noetherian module, +with `f (n+1)` disjoint from the supremum of `f 0`, ..., `f n`, +is eventually zero. -/ +theorem IsNoetherian.disjoint_partialSups_eventually_bot + (f : ℕ → Submodule R M) (h : ∀ n, Disjoint (partialSups f n) (f (n + 1))) : + ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊥ := by + -- A little off-by-one cleanup first: + suffices t : ∃ n : ℕ, ∀ m, n ≤ m → f (m + 1) = ⊥ by + obtain ⟨n, w⟩ := t + use n + 1 + rintro (_ | m) p + · cases p + · apply w + exact Nat.succ_le_succ_iff.mp p + obtain ⟨n, w⟩ := monotone_stabilizes_iff_noetherian.mpr inferInstance (partialSups f) + exact + ⟨n, fun m p => + (h m).eq_bot_of_ge <| sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p⟩ + +end + +-- see Note [lower instance priority] +/-- Modules over the trivial ring are Noetherian. -/ +instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] + [AddCommMonoid M] [Module R M] : IsNoetherian R M := + haveI := Module.subsingleton R M + isNoetherian_of_finite R M + +theorem isNoetherian_of_submodule_of_noetherian (R M) [Semiring R] [AddCommMonoid M] [Module R M] + (N : Submodule R M) (h : IsNoetherian R M) : IsNoetherian R N := + isNoetherian_mk ⟨OrderEmbedding.wellFounded (Submodule.MapSubtype.orderEmbedding N).dual h.wf⟩ + +/-- If `M / S / R` is a scalar tower, and `M / R` is Noetherian, then `M / S` is +also noetherian. -/ +theorem isNoetherian_of_tower (R) {S M} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] + [Module S M] [Module R M] [IsScalarTower R S M] (h : IsNoetherian R M) : IsNoetherian S M := + isNoetherian_mk ⟨(Submodule.restrictScalarsEmbedding R S M).dual.wellFounded h.wf⟩ + +theorem isNoetherian_of_fg_of_noetherian {R M} [Ring R] [AddCommGroup M] [Module R M] + (N : Submodule R M) [I : IsNoetherianRing R] (hN : N.FG) : IsNoetherian R N := by + let ⟨s, hs⟩ := hN + haveI := Classical.decEq M + haveI := Classical.decEq R + have : ∀ x ∈ s, x ∈ N := fun x hx => hs ▸ Submodule.subset_span hx + refine + @isNoetherian_of_surjective + R ((↑s : Set M) → R) N _ _ _ (Pi.module _ _ _) _ ?_ ?_ isNoetherian_pi + · fapply LinearMap.mk + · fapply AddHom.mk + · exact fun f => ⟨∑ i ∈ s.attach, f i • i.1, N.sum_mem fun c _ => N.smul_mem _ <| this _ c.2⟩ + · intro f g + apply Subtype.eq + change (∑ i ∈ s.attach, (f i + g i) • _) = _ + simp only [add_smul, Finset.sum_add_distrib] + rfl + · intro c f + apply Subtype.eq + change (∑ i ∈ s.attach, (c • f i) • _) = _ + simp only [smul_eq_mul, mul_smul] + exact Finset.smul_sum.symm + · rw [LinearMap.range_eq_top] + rintro ⟨n, hn⟩ + change n ∈ N at hn + rw [← hs, ← Set.image_id (s : Set M), Finsupp.mem_span_image_iff_linearCombination] at hn + rcases hn with ⟨l, hl1, hl2⟩ + refine ⟨fun x => l x, Subtype.ext ?_⟩ + change (∑ i ∈ s.attach, l i • (i : M)) = n + rw [s.sum_attach fun i ↦ l i • i, ← hl2, + Finsupp.linearCombination_apply, Finsupp.sum, eq_comm] + refine Finset.sum_subset hl1 fun x _ hx => ?_ + rw [Finsupp.not_mem_support_iff.1 hx, zero_smul] + +instance isNoetherian_of_isNoetherianRing_of_finite (R M : Type*) + [Ring R] [AddCommGroup M] [Module R M] [IsNoetherianRing R] [Module.Finite R M] : + IsNoetherian R M := + have : IsNoetherian R (⊤ : Submodule R M) := + isNoetherian_of_fg_of_noetherian _ <| Module.finite_def.mp inferInstance + isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) + +/-- In a module over a Noetherian ring, the submodule generated by finitely many vectors is +Noetherian. -/ +theorem isNoetherian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R M] + [IsNoetherianRing R] {A : Set M} (hA : A.Finite) : IsNoetherian R (Submodule.span R A) := + isNoetherian_of_fg_of_noetherian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) + +theorem isNoetherianRing_of_surjective (R) [Ring R] (S) [Ring S] (f : R →+* S) + (hf : Function.Surjective f) [H : IsNoetherianRing R] : IsNoetherianRing S := + isNoetherian_mk ⟨OrderEmbedding.wellFounded (Ideal.orderEmbeddingOfSurjective f hf).dual H.wf⟩ + +instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsNoetherianRing R] : + IsNoetherianRing f.range := + isNoetherianRing_of_surjective R f.range f.rangeRestrict f.rangeRestrict_surjective + +theorem isNoetherianRing_of_ringEquiv (R) [Ring R] {S} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] : + IsNoetherianRing S := + isNoetherianRing_of_surjective R S f.toRingHom f.toEquiv.surjective diff --git a/Mathlib/RingTheory/Noetherian/Defs.lean b/Mathlib/RingTheory/Noetherian/Defs.lean index 926f9ba8c6880..e14e88932c6a3 100644 --- a/Mathlib/RingTheory/Noetherian/Defs.lean +++ b/Mathlib/RingTheory/Noetherian/Defs.lean @@ -3,10 +3,7 @@ Copyright (c) 2018 Mario Carneiro, Kevin Buzzard. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kevin Buzzard -/ -import Mathlib.LinearAlgebra.Quotient.Basic -import Mathlib.RingTheory.Finiteness.Cardinality -import Mathlib.RingTheory.Finiteness.Finsupp -import Mathlib.RingTheory.Ideal.Maps +import Mathlib.RingTheory.Finiteness.Basic /-! # Noetherian rings and modules @@ -49,6 +46,9 @@ Noetherian, noetherian, Noetherian ring, Noetherian module, noetherian ring, noe -/ +assert_not_exists Finsupp.linearCombination +assert_not_exists Matrix +assert_not_exists Pi.basis open Set Pointwise @@ -100,144 +100,8 @@ theorem isNoetherian_of_le {s t : Submodule R M} [ht : IsNoetherian R t] (h : s IsNoetherian R s := isNoetherian_submodule.mpr fun _ hs' => isNoetherian_submodule.mp ht _ (le_trans hs' h) -variable (M) - -theorem isNoetherian_of_surjective (f : M →ₗ[R] P) (hf : LinearMap.range f = ⊤) [IsNoetherian R M] : - IsNoetherian R P := - ⟨fun s => - have : (s.comap f).map f = s := Submodule.map_comap_eq_self <| hf.symm ▸ le_top - this ▸ (noetherian _).map _⟩ - -variable {M} - -instance isNoetherian_range (f : M →ₗ[R] P) [IsNoetherian R M] : - IsNoetherian R (LinearMap.range f) := - isNoetherian_of_surjective _ _ f.range_rangeRestrict - -instance isNoetherian_quotient {A M : Type*} [Ring A] [AddCommGroup M] [SMul R A] [Module R M] - [Module A M] [IsScalarTower R A M] (N : Submodule A M) [IsNoetherian R M] : - IsNoetherian R (M ⧸ N) := - isNoetherian_of_surjective M ((Submodule.mkQ N).restrictScalars R) <| - LinearMap.range_eq_top.mpr N.mkQ_surjective - -@[deprecated (since := "2024-04-27"), nolint defLemma] -alias Submodule.Quotient.isNoetherian := isNoetherian_quotient - -theorem isNoetherian_of_linearEquiv (f : M ≃ₗ[R] P) [IsNoetherian R M] : IsNoetherian R P := - isNoetherian_of_surjective _ f.toLinearMap f.range - -theorem LinearEquiv.isNoetherian_iff (f : M ≃ₗ[R] P) : IsNoetherian R M ↔ IsNoetherian R P := - ⟨fun _ ↦ isNoetherian_of_linearEquiv f, fun _ ↦ isNoetherian_of_linearEquiv f.symm⟩ - -theorem isNoetherian_top_iff : IsNoetherian R (⊤ : Submodule R M) ↔ IsNoetherian R M := - Submodule.topEquiv.isNoetherian_iff - -theorem isNoetherian_of_injective [IsNoetherian R P] (f : M →ₗ[R] P) (hf : Function.Injective f) : - IsNoetherian R M := - isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f hf).symm - -theorem fg_of_injective [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) - (hf : Function.Injective f) : N.FG := - haveI := isNoetherian_of_injective f hf - IsNoetherian.noetherian N - -end - -namespace Module - -variable {R M N : Type*} -variable [Semiring R] [AddCommMonoid M] [AddCommMonoid N] [Module R M] [Module R N] -variable (R M) - --- see Note [lower instance priority] -instance (priority := 80) _root_.isNoetherian_of_finite [Finite M] : IsNoetherian R M := - ⟨fun s => ⟨(s : Set M).toFinite.toFinset, by rw [Set.Finite.coe_toFinset, Submodule.span_eq]⟩⟩ - --- see Note [lower instance priority] -instance (priority := 100) IsNoetherian.finite [IsNoetherian R M] : Module.Finite R M := - ⟨IsNoetherian.noetherian ⊤⟩ - -instance {R₁ S : Type*} [CommSemiring R₁] [Semiring S] [Algebra R₁ S] - [IsNoetherian R₁ S] (I : Ideal S) : Module.Finite R₁ I := - IsNoetherian.finite R₁ ((I : Submodule S S).restrictScalars R₁) - -variable {R M} - -theorem Finite.of_injective [IsNoetherian R N] (f : M →ₗ[R] N) (hf : Function.Injective f) : - Module.Finite R M := - ⟨fg_of_injective f hf⟩ - -end Module - -section - -variable {R : Type*} {M : Type*} {P : Type*} -variable [Ring R] [AddCommGroup M] [AddCommGroup P] -variable [Module R M] [Module R P] - -open IsNoetherian - -theorem isNoetherian_of_ker_bot [IsNoetherian R P] (f : M →ₗ[R] P) (hf : LinearMap.ker f = ⊥) : - IsNoetherian R M := - isNoetherian_of_linearEquiv (LinearEquiv.ofInjective f <| LinearMap.ker_eq_bot.mp hf).symm - -theorem fg_of_ker_bot [IsNoetherian R P] {N : Submodule R M} (f : M →ₗ[R] P) - (hf : LinearMap.ker f = ⊥) : N.FG := - haveI := isNoetherian_of_ker_bot f hf - IsNoetherian.noetherian N - -instance isNoetherian_prod [IsNoetherian R M] [IsNoetherian R P] : IsNoetherian R (M × P) := - ⟨fun s => - Submodule.fg_of_fg_map_of_fg_inf_ker (LinearMap.snd R M P) (noetherian _) <| - have : s ⊓ LinearMap.ker (LinearMap.snd R M P) ≤ LinearMap.range (LinearMap.inl R M P) := - fun x ⟨_, hx2⟩ => ⟨x.1, Prod.ext rfl <| Eq.symm <| LinearMap.mem_ker.1 hx2⟩ - Submodule.map_comap_eq_self this ▸ (noetherian _).map _⟩ - -instance isNoetherian_sup (M₁ M₂ : Submodule R P) [IsNoetherian R M₁] [IsNoetherian R M₂] : - IsNoetherian R ↥(M₁ ⊔ M₂) := by - have := isNoetherian_range (M₁.subtype.coprod M₂.subtype) - rwa [LinearMap.range_coprod, Submodule.range_subtype, Submodule.range_subtype] at this - -variable {ι : Type*} [Finite ι] - -instance isNoetherian_pi : - ∀ {M : ι → Type*} [∀ i, AddCommGroup (M i)] - [∀ i, Module R (M i)] [∀ i, IsNoetherian R (M i)], IsNoetherian R (∀ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · exact fun e h ↦ isNoetherian_of_linearEquiv (LinearEquiv.piCongrLeft R _ e) - · infer_instance - · exact fun ih ↦ isNoetherian_of_linearEquiv (LinearEquiv.piOptionEquivProd R).symm - -/-- A version of `isNoetherian_pi` for non-dependent functions. We need this instance because -sometimes Lean fails to apply the dependent version in non-dependent settings (e.g., it fails to -prove that `ι → ℝ` is finite dimensional over `ℝ`). -/ -instance isNoetherian_pi' [IsNoetherian R M] : IsNoetherian R (ι → M) := - isNoetherian_pi - -instance isNoetherian_iSup : - ∀ {M : ι → Submodule R P} [∀ i, IsNoetherian R (M i)], IsNoetherian R ↥(⨆ i, M i) := by - apply Finite.induction_empty_option _ _ _ ι - · intro _ _ e h _ _; rw [← e.iSup_comp]; apply h - · intros; rw [iSup_of_empty]; infer_instance - · intro _ _ ih _ _; rw [iSup_option]; infer_instance - end -section CommRing - -variable (R M N : Type*) [CommRing R] [AddCommGroup M] [AddCommGroup N] [Module R M] [Module R N] - [IsNoetherian R M] [Module.Finite R N] - -instance isNoetherian_linearMap_pi {ι : Type*} [Finite ι] : IsNoetherian R ((ι → R) →ₗ[R] M) := - let _i : Fintype ι := Fintype.ofFinite ι; isNoetherian_of_linearEquiv (Module.piEquiv ι R M) - -instance isNoetherian_linearMap : IsNoetherian R (N →ₗ[R] M) := by - obtain ⟨n, f, hf⟩ := Module.Finite.exists_fin' R N - let g : (N →ₗ[R] M) →ₗ[R] (Fin n → R) →ₗ[R] M := (LinearMap.llcomp R (Fin n → R) N M).flip f - exact isNoetherian_of_injective g hf.injective_linearMapComp_right - -end CommRing - open IsNoetherian Submodule Function section @@ -299,84 +163,6 @@ theorem monotone_stabilizes_iff_noetherian : (∀ f : ℕ →o Submodule R M, ∃ n, ∀ m, n ≤ m → f n = f m) ↔ IsNoetherian R M := by rw [isNoetherian_iff, WellFounded.monotone_chain_condition] -/-- If `∀ I > J, P I` implies `P J`, then `P` holds for all submodules. -/ -theorem IsNoetherian.induction [IsNoetherian R M] {P : Submodule R M → Prop} - (hgt : ∀ I, (∀ J > I, P J) → P I) (I : Submodule R M) : P I := - IsWellFounded.induction _ I hgt - -end - -section - -universe w - -variable {R M P : Type*} {N : Type w} [Ring R] [AddCommGroup M] [Module R M] [AddCommGroup N] - [Module R N] [AddCommGroup P] [Module R P] [IsNoetherian R M] - -lemma Submodule.finite_ne_bot_of_iSupIndep {ι : Type*} {N : ι → Submodule R M} - (h : iSupIndep N) : - Set.Finite {i | N i ≠ ⊥} := - WellFoundedGT.finite_ne_bot_of_iSupIndep h - -@[deprecated (since := "2024-11-24")] -alias Submodule.finite_ne_bot_of_independent := Submodule.finite_ne_bot_of_iSupIndep - -/-- A linearly-independent family of vectors in a module over a non-trivial ring must be finite if -the module is Noetherian. -/ -theorem LinearIndependent.finite_of_isNoetherian [Nontrivial R] {ι} {v : ι → M} - (hv : LinearIndependent R v) : Finite ι := by - refine WellFoundedGT.finite_of_iSupIndep - hv.iSupIndep_span_singleton - fun i contra => ?_ - apply hv.ne_zero i - have : v i ∈ R ∙ v i := Submodule.mem_span_singleton_self (v i) - rwa [contra, Submodule.mem_bot] at this - -theorem LinearIndependent.set_finite_of_isNoetherian [Nontrivial R] {s : Set M} - (hi : LinearIndependent R ((↑) : s → M)) : s.Finite := - @Set.toFinite _ _ hi.finite_of_isNoetherian - -/-- If the first and final modules in an exact sequence are Noetherian, - then the middle module is also Noetherian. -/ -theorem isNoetherian_of_range_eq_ker [IsNoetherian R P] - (f : M →ₗ[R] N) (g : N →ₗ[R] P) (h : LinearMap.range f = LinearMap.ker g) : - IsNoetherian R N := - isNoetherian_mk <| - wellFounded_gt_exact_sequence - (LinearMap.range f) - (Submodule.map (f.ker.liftQ f le_rfl)) - (Submodule.comap (f.ker.liftQ f le_rfl)) - (Submodule.comap g.rangeRestrict) (Submodule.map g.rangeRestrict) - (Submodule.gciMapComap <| LinearMap.ker_eq_bot.mp <| Submodule.ker_liftQ_eq_bot _ _ _ le_rfl) - (Submodule.giMapComap g.surjective_rangeRestrict) - (by simp [Submodule.map_comap_eq, inf_comm, Submodule.range_liftQ]) - (by simp [Submodule.comap_map_eq, h]) - -theorem isNoetherian_iff_submodule_quotient (S : Submodule R P) : - IsNoetherian R P ↔ IsNoetherian R S ∧ IsNoetherian R (P ⧸ S) := by - refine ⟨fun _ ↦ ⟨inferInstance, inferInstance⟩, fun ⟨_, _⟩ ↦ ?_⟩ - apply isNoetherian_of_range_eq_ker S.subtype S.mkQ - rw [Submodule.ker_mkQ, Submodule.range_subtype] - -/-- A sequence `f` of submodules of a noetherian module, -with `f (n+1)` disjoint from the supremum of `f 0`, ..., `f n`, -is eventually zero. -/ -theorem IsNoetherian.disjoint_partialSups_eventually_bot - (f : ℕ → Submodule R M) (h : ∀ n, Disjoint (partialSups f n) (f (n + 1))) : - ∃ n : ℕ, ∀ m, n ≤ m → f m = ⊥ := by - -- A little off-by-one cleanup first: - suffices t : ∃ n : ℕ, ∀ m, n ≤ m → f (m + 1) = ⊥ by - obtain ⟨n, w⟩ := t - use n + 1 - rintro (_ | m) p - · cases p - · apply w - exact Nat.succ_le_succ_iff.mp p - obtain ⟨n, w⟩ := monotone_stabilizes_iff_noetherian.mpr inferInstance (partialSups f) - exact - ⟨n, fun m p => - (h m).eq_bot_of_ge <| sup_eq_left.1 <| (w (m + 1) <| le_add_right p).symm.trans <| w m p⟩ - end /-- A (semi)ring is Noetherian if it is Noetherian as a module over itself, @@ -391,79 +177,3 @@ theorem isNoetherianRing_iff {R} [Semiring R] : IsNoetherianRing R ↔ IsNoether theorem isNoetherianRing_iff_ideal_fg (R : Type*) [Semiring R] : IsNoetherianRing R ↔ ∀ I : Ideal R, I.FG := isNoetherianRing_iff.trans isNoetherian_def - --- see Note [lower instance priority] -/-- Modules over the trivial ring are Noetherian. -/ -instance (priority := 100) isNoetherian_of_subsingleton (R M) [Subsingleton R] [Semiring R] - [AddCommMonoid M] [Module R M] : IsNoetherian R M := - haveI := Module.subsingleton R M - isNoetherian_of_finite R M - -theorem isNoetherian_of_submodule_of_noetherian (R M) [Semiring R] [AddCommMonoid M] [Module R M] - (N : Submodule R M) (h : IsNoetherian R M) : IsNoetherian R N := - isNoetherian_mk ⟨OrderEmbedding.wellFounded (Submodule.MapSubtype.orderEmbedding N).dual h.wf⟩ - -/-- If `M / S / R` is a scalar tower, and `M / R` is Noetherian, then `M / S` is -also noetherian. -/ -theorem isNoetherian_of_tower (R) {S M} [Semiring R] [Semiring S] [AddCommMonoid M] [SMul R S] - [Module S M] [Module R M] [IsScalarTower R S M] (h : IsNoetherian R M) : IsNoetherian S M := - isNoetherian_mk ⟨(Submodule.restrictScalarsEmbedding R S M).dual.wellFounded h.wf⟩ - -theorem isNoetherian_of_fg_of_noetherian {R M} [Ring R] [AddCommGroup M] [Module R M] - (N : Submodule R M) [I : IsNoetherianRing R] (hN : N.FG) : IsNoetherian R N := by - let ⟨s, hs⟩ := hN - haveI := Classical.decEq M - haveI := Classical.decEq R - have : ∀ x ∈ s, x ∈ N := fun x hx => hs ▸ Submodule.subset_span hx - refine - @isNoetherian_of_surjective - R ((↑s : Set M) → R) N _ _ _ (Pi.module _ _ _) _ ?_ ?_ isNoetherian_pi - · fapply LinearMap.mk - · fapply AddHom.mk - · exact fun f => ⟨∑ i ∈ s.attach, f i • i.1, N.sum_mem fun c _ => N.smul_mem _ <| this _ c.2⟩ - · intro f g - apply Subtype.eq - change (∑ i ∈ s.attach, (f i + g i) • _) = _ - simp only [add_smul, Finset.sum_add_distrib] - rfl - · intro c f - apply Subtype.eq - change (∑ i ∈ s.attach, (c • f i) • _) = _ - simp only [smul_eq_mul, mul_smul] - exact Finset.smul_sum.symm - · rw [LinearMap.range_eq_top] - rintro ⟨n, hn⟩ - change n ∈ N at hn - rw [← hs, ← Set.image_id (s : Set M), Finsupp.mem_span_image_iff_linearCombination] at hn - rcases hn with ⟨l, hl1, hl2⟩ - refine ⟨fun x => l x, Subtype.ext ?_⟩ - change (∑ i ∈ s.attach, l i • (i : M)) = n - rw [s.sum_attach fun i ↦ l i • i, ← hl2, - Finsupp.linearCombination_apply, Finsupp.sum, eq_comm] - refine Finset.sum_subset hl1 fun x _ hx => ?_ - rw [Finsupp.not_mem_support_iff.1 hx, zero_smul] - -instance isNoetherian_of_isNoetherianRing_of_finite (R M : Type*) - [Ring R] [AddCommGroup M] [Module R M] [IsNoetherianRing R] [Module.Finite R M] : - IsNoetherian R M := - have : IsNoetherian R (⊤ : Submodule R M) := - isNoetherian_of_fg_of_noetherian _ <| Module.finite_def.mp inferInstance - isNoetherian_of_linearEquiv (LinearEquiv.ofTop (⊤ : Submodule R M) rfl) - -/-- In a module over a Noetherian ring, the submodule generated by finitely many vectors is -Noetherian. -/ -theorem isNoetherian_span_of_finite (R) {M} [Ring R] [AddCommGroup M] [Module R M] - [IsNoetherianRing R] {A : Set M} (hA : A.Finite) : IsNoetherian R (Submodule.span R A) := - isNoetherian_of_fg_of_noetherian _ (Submodule.fg_def.mpr ⟨A, hA, rfl⟩) - -theorem isNoetherianRing_of_surjective (R) [Ring R] (S) [Ring S] (f : R →+* S) - (hf : Function.Surjective f) [H : IsNoetherianRing R] : IsNoetherianRing S := - isNoetherian_mk ⟨OrderEmbedding.wellFounded (Ideal.orderEmbeddingOfSurjective f hf).dual H.wf⟩ - -instance isNoetherianRing_range {R} [Ring R] {S} [Ring S] (f : R →+* S) [IsNoetherianRing R] : - IsNoetherianRing f.range := - isNoetherianRing_of_surjective R f.range f.rangeRestrict f.rangeRestrict_surjective - -theorem isNoetherianRing_of_ringEquiv (R) [Ring R] {S} [Ring S] (f : R ≃+* S) [IsNoetherianRing R] : - IsNoetherianRing S := - isNoetherianRing_of_surjective R S f.toRingHom f.toEquiv.surjective diff --git a/Mathlib/RingTheory/Noetherian/Orzech.lean b/Mathlib/RingTheory/Noetherian/Orzech.lean index 8858b55202f80..da0b31cf04f08 100644 --- a/Mathlib/RingTheory/Noetherian/Orzech.lean +++ b/Mathlib/RingTheory/Noetherian/Orzech.lean @@ -6,7 +6,7 @@ Authors: Mario Carneiro, Kevin Buzzard import Mathlib.Algebra.Module.Submodule.IterateMapComap import Mathlib.Algebra.Order.Archimedean.Basic import Mathlib.Order.PartialSups -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic import Mathlib.RingTheory.OrzechProperty import Mathlib.Order.Filter.AtTopBot diff --git a/Mathlib/RingTheory/Polynomial/Basic.lean b/Mathlib/RingTheory/Polynomial/Basic.lean index 99c10175f0d49..4bee44c7de8c5 100644 --- a/Mathlib/RingTheory/Polynomial/Basic.lean +++ b/Mathlib/RingTheory/Polynomial/Basic.lean @@ -8,7 +8,7 @@ import Mathlib.Algebra.GeomSum import Mathlib.Algebra.MvPolynomial.CommRing import Mathlib.Algebra.MvPolynomial.Equiv import Mathlib.Algebra.Polynomial.BigOperators -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Ring-theoretic supplement of Algebra.Polynomial. diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index 29c25b19c2cf9..dd7f6aa053423 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -7,7 +7,7 @@ import Mathlib.LinearAlgebra.Finsupp.SumProd import Mathlib.RingTheory.Ideal.Prod import Mathlib.RingTheory.Localization.Ideal import Mathlib.RingTheory.Nilpotent.Lemmas -import Mathlib.RingTheory.Noetherian.Defs +import Mathlib.RingTheory.Noetherian.Basic /-! # Prime spectrum of a commutative (semi)ring as a type diff --git a/Mathlib/RingTheory/PrincipalIdealDomain.lean b/Mathlib/RingTheory/PrincipalIdealDomain.lean index 3a675457d9104..5343a17c2253a 100644 --- a/Mathlib/RingTheory/PrincipalIdealDomain.lean +++ b/Mathlib/RingTheory/PrincipalIdealDomain.lean @@ -5,6 +5,7 @@ Authors: Chris Hughes, Morenikeji Neri -/ import Mathlib.Algebra.EuclideanDomain.Field import Mathlib.Algebra.GCDMonoid.Basic +import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Ideal.Nonunits import Mathlib.RingTheory.Noetherian.UniqueFactorizationDomain From 3fa5f452ee1652a47dcf496bc3d52db2f2bdd810 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 13:34:38 +0000 Subject: [PATCH 144/250] chore: ensure there is a replacement identifier or text suggestion for all deprecations (#19426) This is for compatibility with the stricter requirements of leanprover/lean4#6112. Anyone is welcome to replace the "No deprecation message was provided." messages in other PRs! :-) --- Mathlib/Algebra/BigOperators/Group/List.lean | 2 +- Mathlib/Algebra/Group/Action/Defs.lean | 8 +- Mathlib/Algebra/Group/AddChar.lean | 2 +- Mathlib/Algebra/Group/Equiv/Basic.lean | 20 +++- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 29 +++++- Mathlib/Algebra/Module/LinearMap/Defs.lean | 4 +- Mathlib/Algebra/Order/Field/Power.lean | 4 +- .../Algebra/Order/Group/Pointwise/Bounds.lean | 7 +- Mathlib/Analysis/Complex/Circle.lean | 6 +- Mathlib/CategoryTheory/Adjunction/Limits.lean | 8 +- .../CategoryTheory/Idempotents/Karoubi.lean | 2 +- Mathlib/CategoryTheory/Limits/Creates.lean | 12 +-- .../Limits/FunctorCategory/Basic.lean | 8 +- Mathlib/Data/Complex/Abs.lean | 2 +- Mathlib/Data/Fin/Basic.lean | 2 +- Mathlib/Data/List/Basic.lean | 8 +- Mathlib/Data/List/Flatten.lean | 3 +- Mathlib/Data/List/Indexes.lean | 12 +-- Mathlib/Data/List/Lex.lean | 2 +- Mathlib/Data/List/Pairwise.lean | 3 +- Mathlib/Data/List/Permutation.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 9 +- Mathlib/Data/Setoid/Basic.lean | 4 +- Mathlib/Deprecated/AlgebraClasses.lean | 47 ++++----- Mathlib/Deprecated/ByteArray.lean | 10 +- Mathlib/Deprecated/Combinator.lean | 9 +- Mathlib/Deprecated/Equiv.lean | 4 +- Mathlib/Deprecated/LazyList.lean | 10 +- Mathlib/Deprecated/Logic.lean | 74 +++++++-------- Mathlib/Deprecated/NatLemmas.lean | 8 +- Mathlib/Deprecated/RelClasses.lean | 10 +- Mathlib/GroupTheory/Coset/Basic.lean | 8 +- Mathlib/GroupTheory/Coset/Defs.lean | 20 +++- Mathlib/GroupTheory/Exponent.lean | 3 +- Mathlib/GroupTheory/GroupAction/Defs.lean | 8 +- Mathlib/GroupTheory/GroupAction/Quotient.lean | 17 +++- Mathlib/GroupTheory/OrderOfElement.lean | 7 +- .../GroupTheory/SpecificGroups/Cyclic.lean | 9 +- Mathlib/LinearAlgebra/BilinearForm/Basic.lean | 10 +- Mathlib/LinearAlgebra/BilinearForm/Hom.lean | 20 ++-- Mathlib/LinearAlgebra/Prod.lean | 28 +++--- Mathlib/Logic/Basic.lean | 9 +- Mathlib/MeasureTheory/Function/L1Space.lean | 2 +- .../MeasureTheory/Measure/Typeclasses.lean | 2 +- .../Measure/WithDensityFinite.lean | 2 +- Mathlib/NumberTheory/ArithmeticFunction.lean | 3 +- Mathlib/NumberTheory/MulChar/Basic.lean | 6 +- Mathlib/Order/Defs.lean | 8 +- Mathlib/Order/OmegaCompletePartialOrder.lean | 2 +- Mathlib/Order/RelClasses.lean | 6 +- Mathlib/SetTheory/Cardinal/Aleph.lean | 18 ++-- Mathlib/SetTheory/Cardinal/Basic.lean | 2 +- Mathlib/SetTheory/Cardinal/Cofinality.lean | 2 +- Mathlib/SetTheory/Game/Ordinal.lean | 2 +- Mathlib/SetTheory/Ordinal/Arithmetic.lean | 95 ++++++++++--------- Mathlib/SetTheory/Ordinal/Basic.lean | 8 +- Mathlib/SetTheory/Ordinal/Enum.lean | 2 +- Mathlib/SetTheory/Ordinal/FixedPoint.lean | 66 ++++++------- Mathlib/SetTheory/Ordinal/Principal.lean | 4 +- Mathlib/SetTheory/ZFC/Basic.lean | 43 +++++---- Mathlib/Topology/Algebra/ConstMulAction.lean | 2 +- .../Topology/Algebra/ContinuousMonoidHom.lean | 3 +- Mathlib/Topology/Algebra/Group/Quotient.lean | 6 +- Mathlib/Topology/LocalAtTarget.lean | 4 +- Mathlib/Topology/MetricSpace/Polish.lean | 2 +- Mathlib/Topology/NoetherianSpace.lean | 2 +- Mathlib/Topology/Order/OrderClosed.lean | 2 +- 68 files changed, 442 insertions(+), 324 deletions(-) diff --git a/Mathlib/Algebra/BigOperators/Group/List.lean b/Mathlib/Algebra/BigOperators/Group/List.lean index bf01489af4f70..2bd6c8a351d66 100644 --- a/Mathlib/Algebra/BigOperators/Group/List.lean +++ b/Mathlib/Algebra/BigOperators/Group/List.lean @@ -634,7 +634,7 @@ end MonoidHom end MonoidHom set_option linter.deprecated false in -@[simp, deprecated (since := "2024-10-17")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-10-17")] lemma Nat.sum_eq_listSum (l : List ℕ) : Nat.sum l = l.sum := rfl namespace List diff --git a/Mathlib/Algebra/Group/Action/Defs.lean b/Mathlib/Algebra/Group/Action/Defs.lean index 823dbf69f3586..294281be04be5 100644 --- a/Mathlib/Algebra/Group/Action/Defs.lean +++ b/Mathlib/Algebra/Group/Action/Defs.lean @@ -297,9 +297,15 @@ lemma smul_mul_smul_comm [Mul α] [Mul β] [SMul α β] [IsScalarTower α β β] (a • b) * (c • d) = (a * c) • (b * d) := by have : SMulCommClass β α β := .symm ..; exact smul_smul_smul_comm a b c d -@[to_additive (attr := deprecated (since := "2024-08-29"))] +@[to_additive] alias smul_mul_smul := smul_mul_smul_comm +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated smul_mul_smul_comm (since := "2024-08-29")] smul_mul_smul +attribute [deprecated vadd_add_vadd_comm (since := "2024-08-29")] vadd_add_vadd + + /-- Note that the `IsScalarTower α β β` and `SMulCommClass α β β` typeclass arguments are usually satisfied by `Algebra α β`. -/ @[to_additive] diff --git a/Mathlib/Algebra/Group/AddChar.lean b/Mathlib/Algebra/Group/AddChar.lean index 4b0c0f3e8d9bd..0279bc188d39a 100644 --- a/Mathlib/Algebra/Group/AddChar.lean +++ b/Mathlib/Algebra/Group/AddChar.lean @@ -253,7 +253,7 @@ lemma ne_one_iff : ψ ≠ 1 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff lemma ne_zero_iff : ψ ≠ 0 ↔ ∃ x, ψ x ≠ 1 := DFunLike.ne_iff /-- An additive character is *nontrivial* if it takes a value `≠ 1`. -/ -@[deprecated (since := "2024-06-06")] +@[deprecated "No deprecation message was provided." (since := "2024-06-06")] def IsNontrivial (ψ : AddChar A M) : Prop := ∃ a : A, ψ a ≠ 1 set_option linter.deprecated false in diff --git a/Mathlib/Algebra/Group/Equiv/Basic.lean b/Mathlib/Algebra/Group/Equiv/Basic.lean index 203ef93e84ead..b0b66218245be 100644 --- a/Mathlib/Algebra/Group/Equiv/Basic.lean +++ b/Mathlib/Algebra/Group/Equiv/Basic.lean @@ -94,12 +94,28 @@ class MulEquivClass (F : Type*) (A B : outParam Type*) [Mul A] [Mul B] [EquivLik /-- Preserves multiplication. -/ map_mul : ∀ (f : F) (a b), f (a * b) = f a * f b -@[to_additive (attr := deprecated (since := "2024-11-10"))] +@[to_additive] alias MulEquivClass.map_eq_one_iff := EmbeddingLike.map_eq_one_iff -@[to_additive (attr := deprecated (since := "2024-11-10"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_eq_one_iff (since := "2024-11-10")] +MulEquivClass.map_eq_one_iff +attribute [deprecated EmbeddingLike.map_eq_zero_iff (since := "2024-11-10")] +AddEquivClass.map_eq_zero_iff + + +@[to_additive] alias MulEquivClass.map_ne_one_iff := EmbeddingLike.map_ne_one_iff +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated EmbeddingLike.map_ne_one_iff (since := "2024-11-10")] +MulEquivClass.map_ne_one_iff +attribute [deprecated EmbeddingLike.map_ne_zero_iff (since := "2024-11-10")] +AddEquivClass.map_ne_zero_iff + + namespace MulEquivClass variable (F) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index 404fdd46e25ad..d31f9773753a2 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -33,7 +33,13 @@ lemma natCard_mul_le : Nat.card (s * t) ≤ Nat.card s * Nat.card t := by refine Cardinal.toNat_le_toNat Cardinal.mk_mul_le ?_ aesop (add simp [Cardinal.mul_lt_aleph0_iff, finite_mul]) -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_mul_le := natCard_mul_le +@[to_additive] alias card_mul_le := natCard_mul_le + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_mul_le (since := "2024-09-30")] card_mul_le +attribute [deprecated natCard_add_le (since := "2024-09-30")] card_add_le + end Mul @@ -48,7 +54,12 @@ lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by rw [← image_inv, Nat.card_image_of_injective inv_injective] -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_inv := natCard_inv +@[to_additive] alias card_inv := natCard_inv + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_inv (since := "2024-09-30")] card_inv +attribute [deprecated natCard_neg (since := "2024-09-30")] card_neg @[to_additive (attr := simp)] lemma encard_inv (s : Set G) : s⁻¹.encard = s.encard := by simp [encard, PartENat.card] @@ -74,7 +85,12 @@ variable [Group G] {s t : Set G} lemma natCard_div_le : Nat.card (s / t) ≤ Nat.card s * Nat.card t := by rw [div_eq_mul_inv, ← natCard_inv t]; exact natCard_mul_le -@[to_additive (attr := deprecated (since := "2024-09-30"))] alias card_div_le := natCard_div_le +@[to_additive] alias card_div_le := natCard_div_le + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated natCard_div_le (since := "2024-09-30")] card_div_le +attribute [deprecated natCard_sub_le (since := "2024-09-30")] card_sub_le variable [MulAction G α] @@ -86,9 +102,14 @@ lemma _root_.Cardinal.mk_smul_set (a : G) (s : Set α) : #↥(a • s) = #s := lemma natCard_smul_set (a : G) (s : Set α) : Nat.card ↥(a • s) = Nat.card s := Nat.card_image_of_injective (MulAction.injective a) _ -@[to_additive (attr := deprecated (since := "2024-09-30"))] +@[to_additive] alias card_smul_set := Cardinal.mk_smul_set +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Cardinal.mk_smul_set (since := "2024-09-30")] card_smul_set +attribute [deprecated Cardinal.mk_vadd_set (since := "2024-09-30")] card_vadd_set + @[to_additive (attr := simp)] lemma encard_smul_set (a : G) (s : Set α) : (a • s).encard = s.encard := by simp [encard, PartENat.card] diff --git a/Mathlib/Algebra/Module/LinearMap/Defs.lean b/Mathlib/Algebra/Module/LinearMap/Defs.lean index c6436036fb7c1..d4c25f58c52fb 100644 --- a/Mathlib/Algebra/Module/LinearMap/Defs.lean +++ b/Mathlib/Algebra/Module/LinearMap/Defs.lean @@ -596,14 +596,14 @@ variable [Semiring R] [Module R M] [Semiring S] [Module S M₂] [Module R M₃] variable {σ : R →+* S} /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toSemilinearMap (fₗ : M →ₑ+[σ.toMonoidHom] M₂) : M →ₛₗ[σ] M₂ := { fₗ with } instance : SemilinearMapClass (M →ₑ+[σ.toMonoidHom] M₂) σ M M₂ where /-- A `DistribMulActionHom` between two modules is a linear map. -/ -@[deprecated (since := "2024-11-08")] +@[deprecated "No deprecation message was provided." (since := "2024-11-08")] def toLinearMap (fₗ : M →+[R] M₃) : M →ₗ[R] M₃ := { fₗ with } diff --git a/Mathlib/Algebra/Order/Field/Power.lean b/Mathlib/Algebra/Order/Field/Power.lean index d363f2170a5c4..9cec9ad4fcc29 100644 --- a/Mathlib/Algebra/Order/Field/Power.lean +++ b/Mathlib/Algebra/Order/Field/Power.lean @@ -73,13 +73,13 @@ theorem zpow_injective (h₀ : 0 < a) (h₁ : a ≠ 1) : Injective (a ^ · : ℤ theorem zpow_inj (h₀ : 0 < a) (h₁ : a ≠ 1) : a ^ m = a ^ n ↔ m = n := zpow_right_inj₀ h₀ h₁ -@[deprecated (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_of_min_le {x : α} (hx : 1 ≤ x) {a b c : ℤ} (h : min a b ≤ c) : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) := have : Antitone fun n : ℤ => x ^ (-n) := fun _ _ h => zpow_le_zpow_right₀ hx (neg_le_neg h) (this h).trans_eq this.map_min -@[deprecated (since := "2024-10-08")] +@[deprecated "No deprecation message was provided." (since := "2024-10-08")] theorem zpow_le_max_iff_min_le {x : α} (hx : 1 < x) {a b c : ℤ} : x ^ (-c) ≤ max (x ^ (-a)) (x ^ (-b)) ↔ min a b ≤ c := by simp_rw [le_max_iff, min_le_iff, zpow_le_zpow_iff_right₀ hx, neg_le_neg_iff] diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean index ceb329876b8e6..294f6be7c92df 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Bounds.lean @@ -48,7 +48,12 @@ lemma BddAbove.mul (hs : BddAbove s) (ht : BddAbove t) : BddAbove (s * t) := lemma BddBelow.mul (hs : BddBelow s) (ht : BddBelow t) : BddBelow (s * t) := (Nonempty.mul hs ht).mono (subset_lowerBounds_mul s t) -@[to_additive (attr := deprecated (since := "2024-11-13"))] alias Set.BddAbove.mul := BddAbove.mul +@[to_additive] alias Set.BddAbove.mul := BddAbove.mul + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated BddAbove.mul (since := "2024-11-13")] Set.BddAbove.mul +attribute [deprecated BddAbove.add (since := "2024-11-13")] Set.BddAbove.add @[to_additive] lemma BddAbove.range_mul (hf : BddAbove (range f)) (hg : BddAbove (range g)) : diff --git a/Mathlib/Analysis/Complex/Circle.lean b/Mathlib/Analysis/Complex/Circle.lean index 72e66ae46c1b7..e4a9a2e283278 100644 --- a/Mathlib/Analysis/Complex/Circle.lean +++ b/Mathlib/Analysis/Complex/Circle.lean @@ -41,17 +41,17 @@ open ComplexConjugate /-- The unit circle in `ℂ`, here given the structure of a submonoid of `ℂ`. Please use `Circle` when referring to the circle as a type. -/ -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] def circle : Submonoid ℂ := Submonoid.unitSphere ℂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_abs {z : ℂ} : z ∈ circle ↔ abs z = 1 := mem_sphere_zero_iff_norm set_option linter.deprecated false in -@[deprecated (since := "2024-07-24")] +@[deprecated "No deprecation message was provided." (since := "2024-07-24")] theorem mem_circle_iff_normSq {z : ℂ} : z ∈ circle ↔ normSq z = 1 := by simp [Complex.abs, mem_circle_iff_abs] diff --git a/Mathlib/CategoryTheory/Adjunction/Limits.lean b/Mathlib/CategoryTheory/Adjunction/Limits.lean index dd86591976534..2db2c795ef2c2 100644 --- a/Mathlib/CategoryTheory/Adjunction/Limits.lean +++ b/Mathlib/CategoryTheory/Adjunction/Limits.lean @@ -92,7 +92,7 @@ lemma leftAdjoint_preservesColimits : PreservesColimitsOfSize.{v, u} F where ((adj.functorialityAdjunction _).homEquiv _ _)⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma leftAdjointPreservesColimits : PreservesColimitsOfSize.{v, u} F := adj.leftAdjoint_preservesColimits @@ -117,7 +117,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isColimitOfPreserves E.inv t).mapCoconeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsColimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsColimitsOfSize.{v, u} E := Functor.reflectsColimits_of_isEquivalence E @@ -216,7 +216,7 @@ lemma rightAdjoint_preservesLimits : PreservesLimitsOfSize.{v, u} G where ((adj.functorialityAdjunction' _).homEquiv _ _).symm⟩ } } include adj in -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma rightAdjointPreservesLimits : PreservesLimitsOfSize.{v, u} G := adj.rightAdjoint_preservesLimits @@ -240,7 +240,7 @@ noncomputable instance (priority := 100) { reflects := fun t => ⟨(isLimitOfPreserves E.inv t).mapConeEquiv E.asEquivalence.unitIso.symm⟩ } } -@[deprecated (since := "2024-11-18")] +@[deprecated "No deprecation message was provided." (since := "2024-11-18")] lemma isEquivalenceReflectsLimits (E : D ⥤ C) [E.IsEquivalence] : ReflectsLimitsOfSize.{v, u} E := Functor.reflectsLimits_of_isEquivalence E diff --git a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean index 1c0f0936e00c6..d4b4fc8a4f69e 100644 --- a/Mathlib/CategoryTheory/Idempotents/Karoubi.lean +++ b/Mathlib/CategoryTheory/Idempotents/Karoubi.lean @@ -111,7 +111,7 @@ theorem comp_f {P Q R : Karoubi C} (f : P ⟶ Q) (g : Q ⟶ R) : (f ≫ g).f = f @[simp] theorem id_f {P : Karoubi C} : Hom.f (𝟙 P) = P.p := rfl -@[deprecated (since := "2024-07-15")] +@[deprecated "No deprecation message was provided." (since := "2024-07-15")] theorem id_eq {P : Karoubi C} : 𝟙 P = ⟨P.p, by repeat' rw [P.idem]⟩ := rfl /-- It is possible to coerce an object of `C` into an object of `Karoubi C`. diff --git a/Mathlib/CategoryTheory/Limits/Creates.lean b/Mathlib/CategoryTheory/Limits/Creates.lean index 148eafbe8efa1..9769ea2520000 100644 --- a/Mathlib/CategoryTheory/Limits/Creates.lean +++ b/Mathlib/CategoryTheory/Limits/Creates.lean @@ -332,7 +332,7 @@ instance (priority := 100) preservesLimit_of_createsLimit_and_hasLimit (K : J ((liftedLimitMapsToOriginal (limit.isLimit _)).symm ≪≫ (Cones.functoriality K F).mapIso ((liftedLimitIsLimit (limit.isLimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 _ _ @@ -342,7 +342,7 @@ lemma preservesLimitOfCreatesLimitAndHasLimit (K : J ⥤ C) (F : C ⥤ D) 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")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitOfShapeOfCreatesLimitsOfShapeAndHasLimitsOfShape (F : C ⥤ D) [CreatesLimitsOfShape J F] [HasLimitsOfShape J D] : PreservesLimitsOfShape J F := @@ -354,7 +354,7 @@ instance (priority := 100) preservesLimits_of_createsLimits_and_hasLimits (F : C [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesLimitsOfCreatesLimitsAndHasLimits (F : C ⥤ D) [CreatesLimitsOfSize.{w, w'} F] [HasLimitsOfSize.{w, w'} D] : PreservesLimitsOfSize.{w, w'} F := @@ -466,7 +466,7 @@ instance (priority := 100) preservesColimit_of_createsColimit_and_hasColimit (K (Cocones.functoriality K F).mapIso ((liftedColimitIsColimit (colimit.isColimit _)).uniqueUpToIso t))⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 _ _ @@ -477,7 +477,7 @@ instance (priority := 100) preservesColimitOfShape_of_createsColimitsOfShape_and (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitOfShapeOfCreatesColimitsOfShapeAndHasColimitsOfShape (F : C ⥤ D) [CreatesColimitsOfShape J F] [HasColimitsOfShape J D] : PreservesColimitsOfShape J F := @@ -489,7 +489,7 @@ instance (priority := 100) preservesColimits_of_createsColimits_and_hasColimits [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F where -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (since := "2024-11-19")] lemma preservesColimitsOfCreatesColimitsAndHasColimits (F : C ⥤ D) [CreatesColimitsOfSize.{w, w'} F] [HasColimitsOfSize.{w, w'} D] : PreservesColimitsOfSize.{w, w'} F := diff --git a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean index 357916b1fe01e..603001e9e6ce6 100644 --- a/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean +++ b/Mathlib/CategoryTheory/Limits/FunctorCategory/Basic.lean @@ -362,7 +362,7 @@ lemma preservesLimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsLimit ((F ⋙ (evaluation K C).obj X).mapCone c) exact isLimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -374,7 +374,7 @@ lemma preservesLimitsOfShape_of_evaluation (F : D ⥤ K ⥤ C) (J : Type*) [Cate PreservesLimitsOfShape J F := ⟨fun {G} => preservesLimit_of_evaluation F G fun _ => PreservesLimitsOfShape.preservesLimit⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -387,7 +387,7 @@ lemma preservesLimits_of_evaluation (F : D ⥤ K ⥤ C) ⟨fun {L} _ => preservesLimitsOfShape_of_evaluation F L fun _ => PreservesLimitsOfSize.preservesLimitsOfShape⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 := @@ -412,7 +412,7 @@ lemma preservesColimit_of_evaluation (F : D ⥤ K ⥤ C) (G : J ⥤ D) change IsColimit ((F ⋙ (evaluation K C).obj X).mapCocone c) exact isColimitOfPreserves _ hc⟩⟩ -@[deprecated (since := "2024-11-19")] +@[deprecated "No deprecation message was provided." (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 diff --git a/Mathlib/Data/Complex/Abs.lean b/Mathlib/Data/Complex/Abs.lean index 99319c9d298b3..18f2f48e1b525 100644 --- a/Mathlib/Data/Complex/Abs.lean +++ b/Mathlib/Data/Complex/Abs.lean @@ -199,7 +199,7 @@ theorem abs_im_div_abs_le_one (z : ℂ) : |z.im / Complex.abs z| ≤ 1 := @[simp, norm_cast] lemma abs_intCast (n : ℤ) : abs n = |↑n| := by rw [← ofReal_intCast, abs_ofReal] -@[deprecated (since := "2024-02-14")] +@[deprecated "No deprecation message was provided." (since := "2024-02-14")] lemma int_cast_abs (n : ℤ) : |↑n| = Complex.abs n := (abs_intCast _).symm theorem normSq_eq_abs (x : ℂ) : normSq x = (Complex.abs x) ^ 2 := by diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 21c9e27353985..4fcd646907769 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -1115,7 +1115,7 @@ lemma succAbove_ne_last {a : Fin (n + 2)} {b : Fin (n + 1)} (ha : a ≠ last _) lemma succAbove_last_apply (i : Fin n) : succAbove (last n) i = castSucc i := by rw [succAbove_last] -@[deprecated (since := "2024-05-30")] +@[deprecated "No deprecation message was provided." (since := "2024-05-30")] lemma succAbove_lt_ge (p : Fin (n + 1)) (i : Fin n) : castSucc i < p ∨ p ≤ castSucc i := Nat.lt_or_ge (castSucc i) p diff --git a/Mathlib/Data/List/Basic.lean b/Mathlib/Data/List/Basic.lean index 4e3ce8533a890..3f51c9f665d12 100644 --- a/Mathlib/Data/List/Basic.lean +++ b/Mathlib/Data/List/Basic.lean @@ -34,7 +34,7 @@ universe u v w variable {ι : Type*} {α : Type u} {β : Type v} {γ : Type w} {l₁ l₂ : List α} /-- `≤` implies not `>` for lists. -/ -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem le_eq_not_gt [LT α] : ∀ l₁ l₂ : List α, (l₁ ≤ l₂) = ¬l₂ < l₁ := fun _ _ => rfl -- Porting note: Delete this attribute @@ -496,7 +496,7 @@ theorem get_tail (l : List α) (i) (h : i < l.tail.length) l.tail.get ⟨i, h⟩ = l.get ⟨i + 1, h'⟩ := by cases l <;> [cases h; rfl] -@[deprecated (since := "2024-08-22")] +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] theorem get_cons {l : List α} {a : α} {n} (hl) : (a :: l).get ⟨n, hl⟩ = if hn : n = 0 then a else l.get ⟨n - 1, by contrapose! hl; rw [length_cons]; omega⟩ := @@ -627,7 +627,7 @@ lemma cons_sublist_cons' {a b : α} : a :: l₁ <+ b :: l₂ ↔ a :: l₁ <+ l theorem sublist_cons_of_sublist (a : α) (h : l₁ <+ l₂) : l₁ <+ a :: l₂ := h.cons _ -@[deprecated (since := "2024-04-07")] +@[deprecated "No deprecation message was provided." (since := "2024-04-07")] theorem sublist_of_cons_sublist_cons {a} (h : a :: l₁ <+ a :: l₂) : l₁ <+ l₂ := h.of_cons_cons @[deprecated (since := "2024-04-07")] alias cons_sublist_cons_iff := cons_sublist_cons @@ -2281,7 +2281,7 @@ theorem length_dropSlice_lt (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs simp; omega set_option linter.deprecated false in -@[deprecated (since := "2024-07-25")] +@[deprecated "No deprecation message was provided." (since := "2024-07-25")] theorem sizeOf_dropSlice_lt [SizeOf α] (i j : ℕ) (hj : 0 < j) (xs : List α) (hi : i < xs.length) : SizeOf.sizeOf (List.dropSlice i j xs) < SizeOf.sizeOf xs := by induction xs generalizing i j hj with diff --git a/Mathlib/Data/List/Flatten.lean b/Mathlib/Data/List/Flatten.lean index f6fbc8e3044fc..797ba9dc5c297 100644 --- a/Mathlib/Data/List/Flatten.lean +++ b/Mathlib/Data/List/Flatten.lean @@ -158,7 +158,8 @@ theorem append_flatten_map_append (L : List (List α)) (x : List α) : @[deprecated (since := "2024-10-15")] alias append_join_map_append := append_flatten_map_append -@[deprecated (since := "2024-08-15")] theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] +theorem sublist_join {l} {L : List (List α)} (h : l ∈ L) : l <+ L.flatten := sublist_flatten_of_mem h diff --git a/Mathlib/Data/List/Indexes.lean b/Mathlib/Data/List/Indexes.lean index a1b935a65e0cc..fdbe9168bfa47 100644 --- a/Mathlib/Data/List/Indexes.lean +++ b/Mathlib/Data/List/Indexes.lean @@ -89,7 +89,7 @@ theorem mapIdx_eq_ofFn (l : List α) (f : ℕ → α → β) : section deprecated /-- Lean3 `map_with_index` helper function -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List β | _, [] => [] | k, a :: as => f k a :: List.oldMapIdxCore f (k + 1) as @@ -97,12 +97,12 @@ protected def oldMapIdxCore (f : ℕ → α → β) : ℕ → List α → List set_option linter.deprecated false in /-- Given a function `f : ℕ → α → β` and `as : List α`, `as = [a₀, a₁, ...]`, returns the list `[f 0 a₀, f 1 a₁, ...]`. -/ -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected def oldMapIdx (f : ℕ → α → β) (as : List α) : List β := List.oldMapIdxCore f 0 as set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : ℕ) : l.oldMapIdxCore f n = l.oldMapIdx fun i a ↦ f (i + n) a := by induction' l with hd tl hl generalizing f n @@ -111,7 +111,7 @@ protected theorem oldMapIdxCore_eq (l : List α) (f : ℕ → α → β) (n : simp only [List.oldMapIdxCore, hl, Nat.add_left_comm, Nat.add_comm, Nat.add_zero] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) (l₁ l₂ : List α), List.oldMapIdxCore f n (l₁ ++ l₂) = List.oldMapIdxCore f n l₁ ++ List.oldMapIdxCore f (n + l₁.length) l₂ := by @@ -139,7 +139,7 @@ protected theorem oldMapIdxCore_append : ∀ (f : ℕ → α → β) (n : ℕ) ( rw [Nat.add_assoc]; simp only [Nat.add_comm] set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) (e : α), List.oldMapIdx f (l ++ [e]) = List.oldMapIdx f l ++ [f l.length e] := by intros f l e @@ -148,7 +148,7 @@ protected theorem oldMapIdx_append : ∀ (f : ℕ → α → β) (l : List α) ( simp only [Nat.zero_add]; rfl set_option linter.deprecated false in -@[deprecated (since := "2024-08-15")] +@[deprecated "No deprecation message was provided." (since := "2024-08-15")] protected theorem new_def_eq_old_def : ∀ (f : ℕ → α → β) (l : List α), l.mapIdx f = List.oldMapIdx f l := by intro f diff --git a/Mathlib/Data/List/Lex.lean b/Mathlib/Data/List/Lex.lean index fef097d96a69f..69566d9aa5d91 100644 --- a/Mathlib/Data/List/Lex.lean +++ b/Mathlib/Data/List/Lex.lean @@ -99,7 +99,7 @@ instance isAsymm (r : α → α → Prop) [IsAsymm α r] : IsAsymm (List α) (Le | _, _, Lex.cons _, Lex.rel h₂ => asymm h₂ h₂ | _, _, Lex.cons h₁, Lex.cons h₂ => aux _ _ h₁ h₂ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictTotalOrder (r : α → α → Prop) [IsStrictTotalOrder α r] : IsStrictTotalOrder (List α) (Lex r) := { isStrictWeakOrder_of_isOrderConnected with } diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 4682916a610a0..be4f87bb84be6 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -50,7 +50,8 @@ theorem Pairwise.set_pairwise (hl : Pairwise R l) (hr : Symmetric R) : { x | x hl.forall hr -- Porting note: Duplicate of `pairwise_map` but with `f` explicit. -@[deprecated (since := "2024-02-25")] theorem pairwise_map' (f : β → α) : +@[deprecated "No deprecation message was provided." (since := "2024-02-25")] +theorem pairwise_map' (f : β → α) : ∀ {l : List β}, Pairwise R (map f l) ↔ Pairwise (R on f) l | [] => by simp only [map, Pairwise.nil] | b :: l => by diff --git a/Mathlib/Data/List/Permutation.lean b/Mathlib/Data/List/Permutation.lean index 2ce7b4ce525f3..3e0073e3e2049 100644 --- a/Mathlib/Data/List/Permutation.lean +++ b/Mathlib/Data/List/Permutation.lean @@ -419,7 +419,7 @@ theorem length_permutations'Aux (s : List α) (x : α) : · simp · simpa using IH -@[deprecated (since := "2024-06-12")] +@[deprecated "No deprecation message was provided." (since := "2024-06-12")] theorem permutations'Aux_get_zero (s : List α) (x : α) (hn : 0 < length (permutations'Aux x s) := (by simp)) : (permutations'Aux x s).get ⟨0, hn⟩ = x :: s := diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 85ee5070c28b1..20c3db731c333 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -1092,7 +1092,7 @@ lemma sub_mod_eq_zero_of_mod_eq (h : m % k = n % k) : (m - n) % k = 0 := by lemma one_mod_eq_one : ∀ {n : ℕ}, 1 % n = 1 ↔ n ≠ 1 | 0 | 1 | n + 2 => by simp -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] lemma one_mod_of_ne_one : ∀ {n : ℕ}, n ≠ 1 → 1 % n = 1 := one_mod_eq_one.mpr lemma dvd_sub_mod (k : ℕ) : n ∣ k - k % n := diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index 1f4a5da3be5cd..e7008d4c67fbf 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -415,8 +415,13 @@ lemma disjoint_smul_set_left : Disjoint (a • s) t ↔ Disjoint s (a⁻¹ • t lemma disjoint_smul_set_right : Disjoint s (a • t) ↔ Disjoint (a⁻¹ • s) t := by simpa using disjoint_smul_set (a := a) (s := a⁻¹ • s) -@[to_additive (attr := deprecated (since := "2024-10-18"))] -alias smul_set_disjoint_iff := disjoint_smul_set +@[to_additive] alias smul_set_disjoint_iff := disjoint_smul_set + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated disjoint_smul_set (since := "2024-10-18")] smul_set_disjoint_iff +attribute [deprecated disjoint_vadd_set (since := "2024-10-18")] vadd_set_disjoint_iff + end Group diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index 2e75e96025a44..bbaa5b239fbfe 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -36,12 +36,12 @@ attribute [trans] Setoid.trans variable {α : Type*} {β : Type*} /-- A version of `Setoid.r` that takes the equivalence relation as an explicit argument. -/ -@[deprecated (since := "2024-08-29")] +@[deprecated "No deprecation message was provided." (since := "2024-08-29")] def Setoid.Rel (r : Setoid α) : α → α → Prop := @Setoid.r _ r set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] instance Setoid.decidableRel (r : Setoid α) [h : DecidableRel r.r] : DecidableRel r.Rel := h diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index cc26f497422b2..bb6f948871913 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -25,16 +25,16 @@ universe u v variable {α : Sort u} -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsLeftCancel (α : Sort u) (op : α → α → α) : Prop where left_cancel : ∀ a b c, op a b = op a c → b = c -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] class IsRightCancel (α : Sort u) (op : α → α → α) : Prop where right_cancel : ∀ a b c, op a b = op c b → a = c /-- `IsTotalPreorder X r` means that the binary relation `r` on `X` is total and a preorder. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsTotalPreorder (α : Sort u) (r : α → α → Prop) extends IsTrans α r, IsTotal α r : Prop /-- Every total pre-order is a pre-order. -/ @@ -45,11 +45,11 @@ instance (priority := 100) isTotalPreorder_isPreorder (α : Sort u) (r : α → /-- `IsIncompTrans X lt` means that for `lt` a binary relation on `X`, the incomparable relation `fun a b => ¬ lt a b ∧ ¬ lt b a` is transitive. -/ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] class IsIncompTrans (α : Sort u) (lt : α → α → Prop) : Prop where incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) (α : Sort u) (lt : α → α → Prop) [IsStrictWeakOrder α lt] : IsIncompTrans α lt := { ‹IsStrictWeakOrder α lt› with } @@ -59,7 +59,7 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := IsIncompTrans.incomp_trans _ _ _ @@ -68,7 +68,8 @@ section ExplicitRelationVariants variable (r) -@[elab_without_expected_type, deprecated (since := "2024-07-30")] +@[elab_without_expected_type, + deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_trans_of [IsIncompTrans α r] {a b c : α} : ¬a ≺ b ∧ ¬b ≺ a → ¬b ≺ c ∧ ¬c ≺ b → ¬a ≺ c ∧ ¬c ≺ a := incomp_trans @@ -85,32 +86,32 @@ variable {r : α → α → Prop} local infixl:50 " ≺ " => r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] def Equiv (a b : α) : Prop := ¬a ≺ b ∧ ¬b ≺ a local infixl:50 " ≈ " => @Equiv _ r -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem esymm {a b : α} : a ≈ b → b ≈ a := fun ⟨h₁, h₂⟩ => ⟨h₂, h₁⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv {a b : α} : a ≈ b → ¬a ≺ b := fun h => h.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem not_lt_of_equiv' {a b : α} : a ≈ b → ¬b ≺ a := fun h => h.2 variable [IsStrictWeakOrder α r] -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem erefl (a : α) : a ≈ a := ⟨irrefl a, irrefl a⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem etrans {a b c : α} : a ≈ b → b ≈ c → a ≈ c := incomp_trans -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isEquiv : IsEquiv α (@Equiv _ r) where refl := erefl trans _ _ _ := etrans @@ -123,7 +124,7 @@ notation:50 a " ≈[" lt "]" b:50 => @Equiv _ lt a b end StrictWeakOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem isStrictWeakOrder_of_isTotalPreorder {α : Sort u} {le : α → α → Prop} {lt : α → α → Prop} [DecidableRel le] [IsTotalPreorder α le] (h : ∀ a b, lt a b ↔ ¬le b a) : IsStrictWeakOrder α lt := @@ -149,20 +150,20 @@ section LinearOrder variable {α : Type*} [LinearOrder α] set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance : IsTotalPreorder α (· ≤ ·) where trans := @le_trans _ _ total := le_total set_option linter.deprecated false in -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance isStrictWeakOrder_of_linearOrder : IsStrictWeakOrder α (· < ·) := have : IsTotalPreorder α (· ≤ ·) := by infer_instance -- Porting note: added isStrictWeakOrder_of_isTotalPreorder lt_iff_not_ge end LinearOrder -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, lt a b → ¬lt b c ∧ ¬lt c b → lt a c := @fun a b c hab ⟨nbc, ncb⟩ => @@ -171,7 +172,7 @@ theorem lt_of_lt_of_incomp {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt a b ∧ ¬lt b a := incomp_trans_of lt ⟨nac, nca⟩ ⟨ncb, nbc⟩ absurd hab this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeakOrder α lt] [DecidableRel lt] : ∀ {a b c}, ¬lt a b ∧ ¬lt b a → lt b c → lt a c := @fun a b c ⟨nab, nba⟩ hbc => @@ -180,7 +181,7 @@ theorem lt_of_incomp_of_lt {α : Sort u} {lt : α → α → Prop} [IsStrictWeak have : ¬lt b c ∧ ¬lt c b := incomp_trans_of lt ⟨nba, nab⟩ ⟨nac, nca⟩ absurd hbc this.1 -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : ¬lt a b ∧ ¬lt b a → a = b := fun ⟨nab, nba⟩ => match trichotomous_of lt a b with @@ -188,17 +189,17 @@ theorem eq_of_incomp {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α | Or.inr (Or.inl hab) => hab | Or.inr (Or.inr hba) => absurd hba nba -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eq_of_eqv_lt {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] {a b} : a ≈[lt]b → a = b := eq_of_incomp -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem incomp_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : ¬lt a b ∧ ¬lt b a ↔ a = b := Iff.intro eq_of_incomp fun hab => hab ▸ And.intro (irrefl_of lt a) (irrefl_of lt a) -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem eqv_lt_iff_eq {α : Sort u} {lt : α → α → Prop} [IsTrichotomous α lt] [IsIrrefl α lt] (a b) : a ≈[lt]b ↔ a = b := incomp_iff_eq a b diff --git a/Mathlib/Deprecated/ByteArray.lean b/Mathlib/Deprecated/ByteArray.lean index c199be41b3f7c..b7146dbf289c6 100644 --- a/Mathlib/Deprecated/ByteArray.lean +++ b/Mathlib/Deprecated/ByteArray.lean @@ -19,7 +19,7 @@ set_option linter.deprecated false namespace Nat /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def Up (ub a i : Nat) := i < a ∧ i < ub theorem Up.next {ub i} (h : i < ub) : Up ub (i+1) i := ⟨Nat.lt_succ_self _, h⟩ @@ -28,13 +28,13 @@ theorem Up.WF (ub) : WellFounded (Up ub) := Subrelation.wf (h₂ := (measure (ub - ·)).wf) fun ⟨ia, iu⟩ ↦ Nat.sub_lt_sub_left iu ia /-- A well-ordered relation for "upwards" induction on the natural numbers up to some bound `ub`. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def upRel (ub : Nat) : WellFoundedRelation Nat := ⟨Up ub, Up.WF ub⟩ end Nat /-- A terminal byte slice, a suffix of a byte array. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] structure ByteSliceT := (arr : ByteArray) (off : Nat) namespace ByteSliceT @@ -66,7 +66,7 @@ def toArray : ByteSlice → ByteArray universe u v /-- The inner loop of the `forIn` implementation for byte slices. -/ -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β → m (ForInStep β)) (arr : ByteArray) (off _end : Nat) (i : Nat) (b : β) : m β := if h : i < _end then do @@ -75,7 +75,7 @@ def forIn.loop {m : Type u → Type v} {β : Type u} [Monad m] (f : UInt8 → β | ForInStep.yield b => have := Nat.Up.next h; loop f arr off _end (i+1) b else pure b -@[deprecated (since := "2024-08-19")] +@[deprecated "No deprecation message was provided." (since := "2024-08-19")] instance {m : Type u → Type v} : ForIn m ByteSlice UInt8 := ⟨fun ⟨arr, off, len⟩ b f ↦ forIn.loop f arr off (off + len) off b⟩ diff --git a/Mathlib/Deprecated/Combinator.lean b/Mathlib/Deprecated/Combinator.lean index 5c79bf0837a59..366fd5997f982 100644 --- a/Mathlib/Deprecated/Combinator.lean +++ b/Mathlib/Deprecated/Combinator.lean @@ -15,8 +15,11 @@ namespace Combinator universe u v w variable {α : Sort u} {β : Sort v} {γ : Sort w} -@[deprecated (since := "2024-07-27")] def I (a : α) := a -@[deprecated (since := "2024-07-27")] def K (a : α) (_b : β) := a -@[deprecated (since := "2024-07-27")] def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def I (a : α) := a +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def K (a : α) (_b : β) := a +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] +def S (x : α → β → γ) (y : α → β) (z : α) := x z (y z) end Combinator diff --git a/Mathlib/Deprecated/Equiv.lean b/Mathlib/Deprecated/Equiv.lean index 74bc9b68ecdd3..b31fe578386cc 100644 --- a/Mathlib/Deprecated/Equiv.lean +++ b/Mathlib/Deprecated/Equiv.lean @@ -21,10 +21,10 @@ variable {α₁ β₁ : Type*} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α set_option linter.deprecated false -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ diff --git a/Mathlib/Deprecated/LazyList.lean b/Mathlib/Deprecated/LazyList.lean index 7a9b820336505..5ba81d38999a8 100644 --- a/Mathlib/Deprecated/LazyList.lean +++ b/Mathlib/Deprecated/LazyList.lean @@ -24,7 +24,7 @@ namespace LazyList open Function /-- Isomorphism between strict and lazy lists. -/ -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] def listEquivLazyList (α : Type*) : List α ≃ LazyList α where toFun := LazyList.ofList invFun := LazyList.toList @@ -39,12 +39,12 @@ def listEquivLazyList (α : Type*) : List α ≃ LazyList α where · simp [toList, ofList] · simpa [ofList, toList] -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : Traversable LazyList where map := @LazyList.traverse Id _ traverse := @LazyList.traverse -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulTraversable LazyList := by apply Equiv.isLawfulTraversable' listEquivLazyList <;> intros <;> ext <;> rename_i f xs · induction xs using LazyList.rec with @@ -71,7 +71,7 @@ instance : LawfulTraversable LazyList := by Function.comp_def, Thunk.pure, ofList] | mk _ ih => apply ih -@[deprecated (since := "2024-07-22"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-07-22"), simp] theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by induction x using LazyList.rec (motive_2 := fun xs => xs.get.bind singleton = xs.get) with | nil => simp [LazyList.bind] @@ -81,7 +81,7 @@ theorem bind_singleton {α} (x : LazyList α) : x.bind singleton = x := by simp [ih] | mk f ih => simp_all -@[deprecated (since := "2024-07-22")] +@[deprecated "No deprecation message was provided." (since := "2024-07-22")] instance : LawfulMonad LazyList := LawfulMonad.mk' (id_map := by intro α xs diff --git a/Mathlib/Deprecated/Logic.lean b/Mathlib/Deprecated/Logic.lean index f88642ec745a7..52d03620dc96a 100644 --- a/Mathlib/Deprecated/Logic.lean +++ b/Mathlib/Deprecated/Logic.lean @@ -37,43 +37,43 @@ local infix:65 (priority := high) " + " => g def Commutative := ∀ a b, a * b = b * a @[deprecated Std.Associative (since := "2024-09-13")] def Associative := ∀ a b c, (a * b) * c = a * (b * c) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftIdentity := ∀ a, one * a = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightIdentity := ∀ a, a * one = a -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightInverse := ∀ a, a * a⁻¹ = one -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftCancelative := ∀ a b c, a * b = a * c → b = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightCancelative := ∀ a b c, a * b = c * b → a = c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def LeftDistributive := ∀ a b c, a * (b + c) = a * b + a * c -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def RightDistributive := ∀ a b c, (a + b) * c = a * c + b * c end Binary @[deprecated (since := "2024-09-03")] alias not_of_eq_false := of_eq_false -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem cast_proof_irrel {β : Sort u} (h₁ h₂ : α = β) (a : α) : cast h₁ a = cast h₂ a := rfl @[deprecated (since := "2024-09-03")] alias eq_rec_heq := eqRec_heq @[deprecated (since := "2024-09-03")] alias heq_prop := proof_irrel_heq -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_left {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a = a') → (h₂ : Eq.rec (motive := fun a _ ↦ φ a) p₁ e = p₂) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem heq_of_eq_rec_right {φ : α → Sort v} {a a' : α} {p₁ : φ a} {p₂ : φ a'} : (e : a' = a) → (h₂ : p₁ = Eq.rec (motive := fun a _ ↦ φ a) p₂ e) → HEq p₁ p₂ | rfl, rfl => HEq.rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem of_heq_true {a : Prop} (h : HEq a True) : a := of_eq_true (eq_of_heq h) -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem eq_rec_compose {α β φ : Sort u} : ∀ (p₁ : β = φ) (p₂ : α = β) (a : α), (Eq.recOn p₁ (Eq.recOn p₂ a : β) : φ) = Eq.recOn (Eq.trans p₂ p₁) a @@ -114,20 +114,20 @@ theorem iff_self_iff (a : Prop) : (a ↔ a) ↔ True := iff_of_eq (iff_self _) /- decidable -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_True' (h : Decidable True) : decide True = true := by simp -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decide_False' (h : Decidable False) : decide False = false := by simp namespace Decidable -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_true [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : p) (h₄ : h₁ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isTrue _ => rfl) h₄ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def recOn_false [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} (h₃ : ¬p) (h₄ : h₂ h₃) : Decidable.recOn h h₂ h₁ := cast (by match h with | .isFalse _ => rfl) h₄ @@ -145,25 +145,25 @@ end Decidable @[deprecated (since := "2024-09-03")] alias decidableTrue := instDecidableTrue @[deprecated (since := "2024-09-03")] alias decidableFalse := instDecidableFalse -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecEq {α : Sort u} (p : α → α → Bool) : Prop := ∀ ⦃x y : α⦄, p x y = true → x = y -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def IsDecRefl {α : Sort u} (p : α → α → Bool) : Prop := ∀ x, p x x = true -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def decidableEq_of_bool_pred {α : Sort u} {p : α → α → Bool} (h₁ : IsDecEq p) (h₂ : IsDecRefl p) : DecidableEq α | x, y => if hp : p x y = true then isTrue (h₁ hp) else isFalse (fun hxy : x = y ↦ absurd (h₂ y) (by rwa [hxy] at hp)) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inl_refl {α : Sort u} [h : DecidableEq α] (a : α) : h a a = isTrue (Eq.refl a) := match h a a with | isTrue _ => rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} (n : a ≠ b) : h a b = isFalse n := match h a b with @@ -171,7 +171,7 @@ theorem decidableEq_inr_neg {α : Sort u} [h : DecidableEq α] {a b : α} /- subsingleton -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h₂ : ¬p → Sort u} [h₃ : ∀ h : p, Subsingleton (h₁ h)] [h₄ : ∀ h : ¬p, Subsingleton (h₂ h)] : Subsingleton (Decidable.recOn h h₂ h₁) := @@ -179,15 +179,15 @@ theorem rec_subsingleton {p : Prop} [h : Decidable p] {h₁ : p → Sort u} {h | isTrue h => h₃ h | isFalse h => h₄ h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_pos {c t e : Prop} [Decidable c] (h : ite c t e) (hc : c) : t := (if_pos hc ▸ h :) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem imp_of_if_neg {c t e : Prop} [Decidable c] (h : ite c t e) (hnc : ¬c) : e := (if_neg hnc ▸ h :) -@[deprecated (since := "2024-09-11")] +@[deprecated "No deprecation message was provided." (since := "2024-09-11")] theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : Decidable c] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -199,7 +199,7 @@ theorem dif_ctx_congr {α : Sort u} {b c : Prop} [dec_b : Decidable b] [dec_c : | isFalse h₁, isTrue h₂ => absurd h₂ (Iff.mp (not_congr h_c) h₁) | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : Decidable c] (h_c : b ↔ c) (h_t : c → (x ↔ u)) (h_e : ¬c → (y ↔ v)) : ite b x y ↔ ite c u v := match dec_b, dec_c with @@ -209,12 +209,12 @@ theorem if_ctx_congr_prop {b c x y u v : Prop} [dec_b : Decidable b] [dec_c : De | isTrue h₁, isFalse h₂ => absurd h₁ (Iff.mpr (not_congr h_c) h₂) -- @[congr] -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_congr_prop {b c x y u v : Prop} [Decidable b] [Decidable c] (h_c : b ↔ c) (h_t : x ↔ u) (h_e : y ↔ v) : ite b x y ↔ ite c u v := if_ctx_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : c → (x ↔ u)) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -222,7 +222,7 @@ theorem if_ctx_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c (h_e : ¬c → (y ↔ v)) : ite b x y ↔ @ite _ c (decidable_of_decidable_of_iff h_c) u v := if_ctx_congr_prop (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h_t : x ↔ u) -- FIXME: after https://github.com/leanprover/lean4/issues/1867 is fixed, -- this should be changed back to: @@ -230,7 +230,7 @@ theorem if_simp_congr_prop {b c x y u v : Prop} [Decidable b] (h_c : b ↔ c) (h (h_e : y ↔ v) : ite b x y ↔ (@ite _ c (decidable_of_decidable_of_iff h_c) u v) := if_ctx_simp_congr_prop h_c (fun _ ↦ h_t) (fun _ ↦ h_e) -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] {x : b → α} {u : c → α} {y : ¬b → α} {v : ¬c → α} (h_c : b ↔ c) (h_t : ∀ h : c, x (Iff.mpr h_c h) = u h) @@ -241,31 +241,31 @@ theorem dif_ctx_simp_congr {α : Sort u} {b c : Prop} [Decidable b] dite b x y = @dite _ c (decidable_of_decidable_of_iff h_c) u v := dif_ctx_congr (dec_c := decidable_of_decidable_of_iff h_c) h_c h_t h_e -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] def AsTrue (c : Prop) [Decidable c] : Prop := if c then True else False -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib def AsFalse (c : Prop) [Decidable c] : Prop := if c then False else True -@[deprecated (since := "2024-09-03")] +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] theorem AsTrue.get {c : Prop} [h₁ : Decidable c] (_ : AsTrue c) : c := match h₁ with | isTrue h_c => h_c /- Equalities for rewriting let-expressions -/ -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_eq {α : Sort u} {β : Sort v} {a₁ a₂ : α} (b : α → β) (h : a₁ = a₂) : (let x : α := a₁; b x) = (let x : α := a₂; b x) := congrArg b h -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_value_heq {α : Sort v} {β : α → Sort u} {a₁ a₂ : α} (b : ∀ x : α, β x) (h : a₁ = a₂) : HEq (let x : α := a₁; b x) (let x : α := a₂; b x) := by cases h; rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_body_eq {α : Sort v} {β : α → Sort u} (a : α) {b₁ b₂ : ∀ x : α, β x} (h : ∀ x, b₁ x = b₂ x) : (let x : α := a; b₁ x) = (let x : α := a; b₂ x) := by exact h _ ▸ rfl -@[deprecated (since := "2024-09-03")] -- unused in Mathlib +@[deprecated "No deprecation message was provided." (since := "2024-09-03")] -- unused in Mathlib theorem let_eq {α : Sort v} {β : Sort u} {a₁ a₂ : α} {b₁ b₂ : α → β} (h₁ : a₁ = a₂) (h₂ : ∀ x, b₁ x = b₂ x) : (let x : α := a₁; b₁ x) = (let x : α := a₂; b₂ x) := by simp [h₁, h₂] diff --git a/Mathlib/Deprecated/NatLemmas.lean b/Mathlib/Deprecated/NatLemmas.lean index 37fd557b3fec3..81b124b4476cd 100644 --- a/Mathlib/Deprecated/NatLemmas.lean +++ b/Mathlib/Deprecated/NatLemmas.lean @@ -28,7 +28,7 @@ namespace Nat /-! successor and predecessor -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ m → B) : B := by induction n with | zero => exact H1 rfl @@ -36,7 +36,7 @@ def discriminate {B : Sort u} {n : ℕ} (H1 : n = 0 → B) (H2 : ∀ m, n = succ -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem one_eq_succ_zero : 1 = succ 0 := rfl @@ -44,7 +44,7 @@ theorem one_eq_succ_zero : 1 = succ 0 := -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P (succ n) 0) (H3 : ∀ n m, P n m → P (succ n) (succ m)) : ∀ n m : ℕ, P n m | 0, _m => H1 _ @@ -55,7 +55,7 @@ def subInduction {P : ℕ → ℕ → Sort u} (H1 : ∀ m, P 0 m) (H2 : ∀ n, P -- Unused in Mathlib; -- if downstream projects find this essential please copy it or remove the deprecation. -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cond_decide_mod_two (x : ℕ) [d : Decidable (x % 2 = 1)] : cond (@decide (x % 2 = 1) d) 1 0 = x % 2 := by simp only [cond_eq_if, decide_eq_true_eq] diff --git a/Mathlib/Deprecated/RelClasses.lean b/Mathlib/Deprecated/RelClasses.lean index 357b10283c2d5..8d5f647030d5d 100644 --- a/Mathlib/Deprecated/RelClasses.lean +++ b/Mathlib/Deprecated/RelClasses.lean @@ -27,12 +27,14 @@ variable {α : Type u} open Function -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where -@[deprecated (since := "2024-08-22")] instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≤ ·) where +@[deprecated "No deprecation message was provided." (since := "2024-08-22")] +instance [LinearOrder α] : IsTotalPreorder α (· ≥ ·) where -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsIncompTrans α (· < ·) := by infer_instance diff --git a/Mathlib/GroupTheory/Coset/Basic.lean b/Mathlib/GroupTheory/Coset/Basic.lean index a93c6f92a2109..4f8c575b67fcf 100644 --- a/Mathlib/GroupTheory/Coset/Basic.lean +++ b/Mathlib/GroupTheory/Coset/Basic.lean @@ -318,9 +318,15 @@ lemma orbit_eq_out_smul (x : α ⧸ s) : MulAction.orbitRel.Quotient.orbit x = x 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"))] +@[to_additive] alias orbit_eq_out'_smul := orbit_eq_out_smul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbit_eq_out_smul (since := "2024-10-19")] orbit_eq_out'_smul +attribute [deprecated QuotientAddGroup.orbit_eq_out_vadd (since := "2024-10-19")] +QuotientAddGroup.orbit_eq_out'_vadd + end QuotientGroup namespace Subgroup diff --git a/Mathlib/GroupTheory/Coset/Defs.lean b/Mathlib/GroupTheory/Coset/Defs.lean index bd49ccf9a7568..c57164adfbaee 100644 --- a/Mathlib/GroupTheory/Coset/Defs.lean +++ b/Mathlib/GroupTheory/Coset/Defs.lean @@ -174,7 +174,13 @@ theorem induction_on {C : α ⧸ s → Prop} (x : α ⧸ s) (H : ∀ z, C (Quoti instance : Coe α (α ⧸ s) := ⟨mk⟩ -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias induction_on' := induction_on +@[to_additive] alias induction_on' := induction_on + +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated induction_on (since := "2024-08-04")] induction_on' +attribute [deprecated QuotientAddGroup.induction_on (since := "2024-08-04")] +QuotientAddGroup.induction_on' @[to_additive (attr := simp)] theorem quotient_liftOn_mk {β} (f : α → β) (h) (x : α) : Quotient.liftOn' (x : α ⧸ s) f h = f x := @@ -198,7 +204,8 @@ protected theorem eq {a b : α} : (a : α ⧸ s) = b ↔ a⁻¹ * b ∈ s := _ ↔ leftRel s a b := Quotient.eq'' _ ↔ _ := by rw [leftRel_apply] -@[to_additive (attr := deprecated (since := "2024-08-04"))] alias eq' := QuotientGroup.eq +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-04"))] +alias eq' := QuotientGroup.eq @[to_additive] theorem out_eq' (a : α ⧸ s) : mk a.out = a := @@ -213,9 +220,16 @@ variable (s) 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] +@[to_additive QuotientAddGroup.mk_out'_eq_mul] alias mk_out'_eq_mul := mk_out_eq_mul +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated mk_out_eq_mul (since := "2024-10-19")] mk_out'_eq_mul +attribute [deprecated QuotientAddGroup.mk_out_eq_mul (since := "2024-10-19")] +QuotientAddGroup.mk_out'_eq_mul + + variable {s} {a b : α} @[to_additive (attr := simp)] diff --git a/Mathlib/GroupTheory/Exponent.lean b/Mathlib/GroupTheory/Exponent.lean index 6ee49e4b8a6ac..ea64a5ce67e5e 100644 --- a/Mathlib/GroupTheory/Exponent.lean +++ b/Mathlib/GroupTheory/Exponent.lean @@ -644,7 +644,8 @@ lemma inv_eq_self_of_orderOf_eq_two {x : G} (hx : orderOf x = 2) : -- TODO: delete /-- Any group of exponent two is abelian. -/ -@[to_additive (attr := reducible, deprecated (since := "2024-02-17")) +@[to_additive (attr := reducible, + deprecated "No deprecation message was provided." (since := "2024-02-17")) "Any additive group of exponent two is abelian."] def instCommGroupOfExponentTwo (hG : Monoid.exponent G = 2) : CommGroup G where mul_comm := mul_comm_of_exponent_two hG diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index 3848050964715..fa15f5e7d6253 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -312,9 +312,15 @@ variable {G α} theorem orbitRel_apply {a b : α} : orbitRel G α a b ↔ a ∈ orbit G b := Iff.rfl -@[to_additive (attr := deprecated (since := "2024-10-18"))] +@[to_additive] alias orbitRel_r_apply := orbitRel_apply +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated orbitRel_apply (since := "2024-10-18")] orbitRel_r_apply +attribute [deprecated AddAction.orbitRel_apply (since := "2024-10-18")] AddAction.orbitRel_r_apply + + /-- When you take a set `U` in `α`, push it down to the quotient, and pull back, you get the union of the orbit of `U` under `G`. -/ @[to_additive diff --git a/Mathlib/GroupTheory/GroupAction/Quotient.lean b/Mathlib/GroupTheory/GroupAction/Quotient.lean index 4c15d5da76105..778fb1159400a 100644 --- a/Mathlib/GroupTheory/GroupAction/Quotient.lean +++ b/Mathlib/GroupTheory/GroupAction/Quotient.lean @@ -103,13 +103,26 @@ theorem _root_.QuotientGroup.out_conj_pow_minimalPeriod_mem (a : α) (q : α ⧸ 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"))] +@[to_additive] alias Quotient.mk_smul_out' := Quotient.mk_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.mk_smul_out (since := "2024-10-19")] Quotient.mk_smul_out' +attribute [deprecated AddAction.Quotient.mk_vadd_out (since := "2024-10-19")] +AddAction.Quotient.mk_vadd_out' + -- Porting note: removed simp attribute, simp can prove this -@[to_additive (attr := deprecated (since := "2024-10-19"))] +@[to_additive] alias Quotient.coe_smul_out' := Quotient.coe_smul_out +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Quotient.coe_smul_out (since := "2024-10-19")] Quotient.coe_smul_out' +attribute [deprecated AddAction.Quotient.coe_vadd_out (since := "2024-10-19")] +AddAction.Quotient.coe_vadd_out' + + @[deprecated (since := "2024-10-19")] alias _root_.QuotientGroup.out'_conj_pow_minimalPeriod_mem := QuotientGroup.out_conj_pow_minimalPeriod_mem diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index b380c1f71c410..b18e30f62e781 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -88,9 +88,14 @@ theorem not_isOfFinOrder_of_injective_pow {x : G} (h : Injective fun n : ℕ => theorem IsOfFinOrder.one : IsOfFinOrder (1 : G) := isOfFinOrder_iff_pow_eq_one.mpr ⟨1, Nat.one_pos, one_pow 1⟩ -@[to_additive (attr := deprecated (since := "2024-10-11"))] +@[to_additive] alias isOfFinOrder_one := IsOfFinOrder.one +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated IsOfFinOrder.one (since := "2024-10-11")] isOfFinOrder_one +attribute [deprecated IsOfFinAddOrder.zero (since := "2024-10-11")] isOfFinAddOrder_zero + @[to_additive] lemma IsOfFinOrder.pow {n : ℕ} : IsOfFinOrder a → IsOfFinOrder (a ^ n) := by simp_rw [isOfFinOrder_iff_pow_eq_one] diff --git a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean index 11e9dc5991b44..d13254a943cc7 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Cyclic.lean @@ -327,10 +327,17 @@ lemma isCyclic_iff_exists_ofOrder_eq_natCard [Finite α] : refine isCyclic_of_orderOf_eq_card g ?_ simp [hg] -@[to_additive (attr := deprecated (since := "2024-04-20"))] +@[to_additive] protected alias IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype := isCyclic_iff_exists_ofOrder_eq_natCard +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated isCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype +attribute [deprecated isAddCyclic_iff_exists_ofOrder_eq_natCard (since := "2024-04-20")] +IsAddCyclic.iff_exists_ofOrder_eq_natCard_of_Fintype + section variable [Fintype α] diff --git a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean index 0768933dedc46..da171b8d6bfaa 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Basic.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Basic.lean @@ -58,7 +58,7 @@ namespace LinearMap namespace BilinForm -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coeFn_congr : ∀ {x x' y y' : M}, x = x' → y = y' → B x y = B x' y' | _, _, _, _, rfl, rfl => rfl @@ -101,7 +101,7 @@ theorem ext (H : ∀ x y : M, B x y = D x y) : B = D := ext₂ H theorem congr_fun (h : B = D) (x y : M) : B x y = D x y := congr_fun₂ h _ _ -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_zero : ⇑(0 : BilinForm R M) = 0 := rfl @@ -111,7 +111,7 @@ theorem zero_apply (x y : M) : (0 : BilinForm R M) x y = 0 := variable (B D B₁ D₁) -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_add : ⇑(B + D) = B + D := rfl @@ -119,7 +119,7 @@ theorem coe_add : ⇑(B + D) = B + D := theorem add_apply (x y : M) : (B + D) x y = B x y + D x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_neg : ⇑(-B₁) = -B₁ := rfl @@ -127,7 +127,7 @@ theorem coe_neg : ⇑(-B₁) = -B₁ := theorem neg_apply (x y : M₁) : (-B₁) x y = -B₁ x y := rfl -@[deprecated (since := "2024-04-14")] +@[deprecated "No deprecation message was provided." (since := "2024-04-14")] theorem coe_sub : ⇑(B₁ - D₁) = B₁ - D₁ := rfl diff --git a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean index 9c011c36ed52b..c1b171f04d13e 100644 --- a/Mathlib/LinearAlgebra/BilinearForm/Hom.lean +++ b/Mathlib/LinearAlgebra/BilinearForm/Hom.lean @@ -58,16 +58,16 @@ section ToLin' def toLinHomAux₁ (A : BilinForm R M) (x : M) : M →ₗ[R] R := A x /-- Auxiliary definition to define `toLinHom`; see below. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHomAux₂ (A : BilinForm R M) : M →ₗ[R] M →ₗ[R] R := A /-- The linear map obtained from a `BilinForm` by fixing the left co-ordinate and evaluating in the right. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def toLinHom : BilinForm R M →ₗ[R] M →ₗ[R] M →ₗ[R] R := LinearMap.id set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem toLin'_apply (A : BilinForm R M) (x : M) : toLinHom (M := M) A x = A x := rfl @@ -112,7 +112,7 @@ def LinearMap.toBilinAux (f : M →ₗ[R] M →ₗ[R] R) : BilinForm R M := f set_option linter.deprecated false in /-- Bilinear forms are linearly equivalent to maps with two arguments that are linear in both. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] R := { BilinForm.toLinHom with invFun := LinearMap.toBilinAux @@ -121,35 +121,35 @@ def LinearMap.BilinForm.toLin : BilinForm R M ≃ₗ[R] M →ₗ[R] M →ₗ[R] set_option linter.deprecated false in /-- A map with two arguments that is linear in both is linearly equivalent to bilinear form. -/ -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] def LinearMap.toBilin : (M →ₗ[R] M →ₗ[R] R) ≃ₗ[R] BilinForm R M := BilinForm.toLin.symm -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilinAux_eq (f : M →ₗ[R] M →ₗ[R] R) : LinearMap.toBilinAux f = f := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_symm : (LinearMap.toBilin.symm : BilinForm R M ≃ₗ[R] _) = BilinForm.toLin := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_symm : (BilinForm.toLin.symm : _ ≃ₗ[R] BilinForm R M) = LinearMap.toBilin := LinearMap.toBilin.symm_symm set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem LinearMap.toBilin_apply (f : M →ₗ[R] M →ₗ[R] R) (x y : M) : toBilin f x y = f x y := rfl set_option linter.deprecated false in -@[deprecated (since := "2024-04-26")] +@[deprecated "No deprecation message was provided." (since := "2024-04-26")] theorem BilinForm.toLin_apply (x : M) : BilinForm.toLin B x = B x := rfl diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 4a19736a68ba4..7bbb4ac2f4b65 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -834,17 +834,17 @@ set_option linter.deprecated false /-- An auxiliary construction for `tunnel`. The composition of `f`, followed by the isomorphism back to `K`, followed by the inclusion of this submodule back into `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnelAux (f : M × N →ₗ[R] M) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : M × N →ₗ[R] M := (Kφ.1.subtype.comp Kφ.2.symm.toLinearMap).comp f -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tunnelAux_injective (f : M × N →ₗ[R] M) (i : Injective f) (Kφ : ΣK : Submodule R M, K ≃ₗ[R] M) : Injective (tunnelAux f Kφ) := (Subtype.val_injective.comp Kφ.2.symm.injective).comp i /-- Auxiliary definition for `tunnel`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule R M, K ≃ₗ[R] M | 0 => ⟨⊤, LinearEquiv.ofTop ⊤ rfl⟩ | n + 1 => @@ -855,7 +855,7 @@ def tunnel' (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → ΣK : Submodule /-- Give an injective map `f : M × N →ₗ[R] M` we can find a nested sequence of submodules all isomorphic to `M`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M)ᵒᵈ := -- Note: the hint `(α := _)` had to be added in https://github.com/leanprover-community/mathlib4/pull/8386 ⟨fun n => OrderDual.toDual (α := Submodule R M) (tunnel' f i n).1, @@ -867,24 +867,24 @@ def tunnel (f : M × N →ₗ[R] M) (i : Injective f) : ℕ →o (Submodule R M) /-- Give an injective map `f : M × N →ₗ[R] M` we can find a sequence of submodules all isomorphic to `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Submodule R M := (Submodule.snd R M N).map (tunnelAux f (tunnel' f i n)) /-- Each `tailing f i n` is a copy of `N`. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailingLinearEquiv (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≃ₗ[R] N := ((Submodule.snd R M N).equivMapOfInjective _ (tunnelAux_injective f i (tunnel' f i n))).symm.trans (Submodule.sndEquiv R M N) -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ≤ OrderDual.ofDual (α := Submodule R M) (tunnel f i n) := by dsimp [tailing, tunnelAux] rw [Submodule.map_comp, Submodule.map_comp] apply Submodule.map_subtype_le -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailing f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by rw [disjoint_iff] @@ -893,7 +893,7 @@ theorem tailing_disjoint_tunnel_succ (f : M × N →ₗ[R] M) (i : Injective f) Submodule.comap_map_eq_of_injective (tunnelAux_injective _ i _), inf_comm, Submodule.fst_inf_snd, Submodule.map_bot] -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailing f i n ⊔ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) ≤ (OrderDual.ofDual (α := Submodule R M) <| tunnel f i n) := by @@ -902,19 +902,19 @@ theorem tailing_sup_tunnel_succ_le_tunnel (f : M × N →ₗ[R] M) (i : Injectiv apply Submodule.map_subtype_le /-- The supremum of all the copies of `N` found inside the tunnel. -/ -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] def tailings (f : M × N →ₗ[R] M) (i : Injective f) : ℕ → Submodule R M := partialSups (tailing f i) -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_zero (f : M × N →ₗ[R] M) (i : Injective f) : tailings f i 0 = tailing f i 0 := by simp [tailings] -@[simp, deprecated (since := "2024-06-05")] +@[simp, deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_succ (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : tailings f i (n + 1) = tailings f i n ⊔ tailing f i (n + 1) := by simp [tailings] -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (OrderDual.ofDual (α := Submodule R M) <| tunnel f i (n + 1)) := by induction' n with n ih @@ -926,7 +926,7 @@ theorem tailings_disjoint_tunnel (f : M × N →ₗ[R] M) (i : Injective f) (n : · apply Disjoint.mono_right _ ih apply tailing_sup_tunnel_succ_le_tunnel -@[deprecated (since := "2024-06-05")] +@[deprecated "No deprecation message was provided." (since := "2024-06-05")] theorem tailings_disjoint_tailing (f : M × N →ₗ[R] M) (i : Injective f) (n : ℕ) : Disjoint (tailings f i n) (tailing f i (n + 1)) := Disjoint.mono_right (tailing_le_tunnel f i _) (tailings_disjoint_tunnel f i _) diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 7a0eb1ff4d84f..89c3a8fe8a2f1 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -725,20 +725,21 @@ alias by_contradiction := byContradiction -- TODO: remove? rename in core? alias prop_complete := propComplete -- TODO: remove? rename in core? -@[elab_as_elim, deprecated (since := "2024-07-27")] theorem cases_true_false (p : Prop → Prop) +@[elab_as_elim, deprecated "No deprecation message was provided." (since := "2024-07-27")] +theorem cases_true_false (p : Prop → Prop) (h1 : p True) (h2 : p False) (a : Prop) : p a := Or.elim (prop_complete a) (fun ht : a = True ↦ ht.symm ▸ h1) fun hf : a = False ↦ hf.symm ▸ h2 -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem eq_false_or_eq_true (a : Prop) : a = False ∨ a = True := (prop_complete a).symm set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p a := @cases_true_false p h1 h2 a set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 end Classical diff --git a/Mathlib/MeasureTheory/Function/L1Space.lean b/Mathlib/MeasureTheory/Function/L1Space.lean index 25e17db9f409b..b003cc9a9db03 100644 --- a/Mathlib/MeasureTheory/Function/L1Space.lean +++ b/Mathlib/MeasureTheory/Function/L1Space.lean @@ -443,7 +443,7 @@ lemma Integrable.of_finite [Finite α] [MeasurableSingletonClass α] [IsFiniteMe /-- This lemma is a special case of `Integrable.of_finite`. -/ -- Eternal deprecation for discoverability, don't remove -@[deprecated Integrable.of_finite, nolint deprecatedNoSince] +@[deprecated Integrable.of_finite (since := "2024-10-05"), nolint deprecatedNoSince] lemma Integrable.of_isEmpty [IsEmpty α] {f : α → β} : Integrable f μ := .of_finite @[deprecated (since := "2024-02-05")] alias integrable_of_fintype := Integrable.of_finite diff --git a/Mathlib/MeasureTheory/Measure/Typeclasses.lean b/Mathlib/MeasureTheory/Measure/Typeclasses.lean index c9a41869344f8..f8d833559d3e9 100644 --- a/Mathlib/MeasureTheory/Measure/Typeclasses.lean +++ b/Mathlib/MeasureTheory/Measure/Typeclasses.lean @@ -567,7 +567,7 @@ instance isFiniteMeasure_sfiniteSeq [h : SFinite μ] (n : ℕ) : IsFiniteMeasure h.1.choose_spec.1 n set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] instance isFiniteMeasure_sFiniteSeq [SFinite μ] (n : ℕ) : IsFiniteMeasure (sFiniteSeq μ n) := isFiniteMeasure_sfiniteSeq n diff --git a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean index 21abd06cd6cee..40b56846328c7 100644 --- a/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean +++ b/Mathlib/MeasureTheory/Measure/WithDensityFinite.lean @@ -117,7 +117,7 @@ noncomputable def Measure.densityToFinite (μ : Measure α) [SFinite μ] (a : α μ.rnDeriv μ.toFinite a set_option linter.deprecated false in -@[deprecated (since := "2024-10-04")] +@[deprecated "No deprecation message was provided." (since := "2024-10-04")] lemma densityToFinite_def (μ : Measure α) [SFinite μ] : μ.densityToFinite = μ.rnDeriv μ.toFinite := rfl diff --git a/Mathlib/NumberTheory/ArithmeticFunction.lean b/Mathlib/NumberTheory/ArithmeticFunction.lean index 5c13f9d53ebe2..8b8038b188ab0 100644 --- a/Mathlib/NumberTheory/ArithmeticFunction.lean +++ b/Mathlib/NumberTheory/ArithmeticFunction.lean @@ -1291,7 +1291,8 @@ theorem _root_.Nat.card_divisors {n : ℕ} (hn : n ≠ 0) : exact Finset.prod_congr n.support_factorization fun _ h => sigma_zero_apply_prime_pow <| Nat.prime_of_mem_primeFactors h -@[deprecated (since := "2024-06-09")] theorem card_divisors (n : ℕ) (hn : n ≠ 0) : +@[deprecated "No deprecation message was provided." (since := "2024-06-09")] +theorem card_divisors (n : ℕ) (hn : n ≠ 0) : #n.divisors = n.primeFactors.prod (n.factorization · + 1) := Nat.card_divisors hn theorem _root_.Nat.sum_divisors {n : ℕ} (hn : n ≠ 0) : diff --git a/Mathlib/NumberTheory/MulChar/Basic.lean b/Mathlib/NumberTheory/MulChar/Basic.lean index 2df970973c59d..7ddd95b920b4e 100644 --- a/Mathlib/NumberTheory/MulChar/Basic.lean +++ b/Mathlib/NumberTheory/MulChar/Basic.lean @@ -391,13 +391,13 @@ lemma ne_one_iff {χ : MulChar R R'} : χ ≠ 1 ↔ ∃ a : Rˣ, χ a ≠ 1 := b simp only [Ne, eq_one_iff, not_forall] /-- A multiplicative character is *nontrivial* if it takes a value `≠ 1` on a unit. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] def IsNontrivial (χ : MulChar R R') : Prop := ∃ a : Rˣ, χ a ≠ 1 set_option linter.deprecated false in /-- A multiplicative character is nontrivial iff it is not the trivial character. -/ -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] theorem isNontrivial_iff (χ : MulChar R R') : χ.IsNontrivial ↔ χ ≠ 1 := by simp only [IsNontrivial, Ne, MulChar.ext_iff, not_forall, one_apply_coe] @@ -573,7 +573,7 @@ theorem sum_eq_zero_of_ne_one [IsDomain R'] {χ : MulChar R R'} (hχ : χ ≠ 1) simpa only [Finset.mul_sum, ← map_mul] using b.mulLeft_bijective.sum_comp _ set_option linter.deprecated false in -@[deprecated (since := "2024-06-16")] +@[deprecated "No deprecation message was provided." (since := "2024-06-16")] lemma IsNontrivial.sum_eq_zero [IsDomain R'] {χ : MulChar R R'} (hχ : χ.IsNontrivial) : ∑ a, χ a = 0 := sum_eq_zero_of_ne_one ((isNontrivial_iff _).mp hχ) diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index dc90cc7b0adbe..b97cdc2c979d6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -206,11 +206,11 @@ theorem Equivalence.transitive (h : Equivalence r) : Transitive r := variable {β : Sort*} (r : β → β → Prop) (f : α → β) -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ -@[deprecated (since := "2024-09-13")] +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ @@ -276,7 +276,7 @@ lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_n lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a | _a, _b, hab => lt_iff_le_not_le.mp hab @@ -502,7 +502,7 @@ namespace Nat /-! Deprecated properties of inequality on `Nat` -/ -@[deprecated (since := "2024-08-23")] +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index 4c97970762171..c34eaa969055e 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -731,7 +731,7 @@ theorem apply_mono {f g : α →𝒄 β} {x y : α} (h₁ : f ≤ g) (h₂ : x OrderHom.apply_mono (show (f : α →o β) ≤ g from h₁) h₂ set_option linter.deprecated false in -@[deprecated (since := "2024-07-27")] +@[deprecated "No deprecation message was provided." (since := "2024-07-27")] theorem ite_continuous' {p : Prop} [hp : Decidable p] (f g : α → β) (hf : Continuous' f) (hg : Continuous' g) : Continuous' fun x => if p then f x else g x := by split_ifs <;> simp [*] diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index c99abc227fd09..0acae0e49466b 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -86,7 +86,7 @@ theorem IsStrictOrder.swap (r) [IsStrictOrder α r] : IsStrictOrder α (swap r) theorem IsPartialOrder.swap (r) [IsPartialOrder α r] : IsPartialOrder α (swap r) := { @IsPreorder.swap α r _, @IsAntisymm.swap α r _ with } -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] theorem IsLinearOrder.swap (r) [IsLinearOrder α r] : IsLinearOrder α (swap r) := { @IsPartialOrder.swap α r _, @IsTotal.swap α r _ with } @@ -212,7 +212,7 @@ instance (priority := 100) isStrictOrderConnected_of_isStrictTotalOrder [IsStric fun o ↦ o.elim (fun e ↦ e ▸ h) fun h' ↦ _root_.trans h' h⟩ -- see Note [lower instance priority] -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance (priority := 100) isStrictTotalOrder_of_isStrictTotalOrder [IsStrictTotalOrder α r] : IsStrictWeakOrder α r := { isStrictWeakOrder_of_isOrderConnected with } @@ -780,7 +780,7 @@ instance [LinearOrder α] : IsStrictTotalOrder α (· < ·) where instance [LinearOrder α] : IsOrderConnected α (· < ·) := by infer_instance -@[deprecated (since := "2024-07-30")] +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] instance [LinearOrder α] : IsStrictWeakOrder α (· < ·) := by infer_instance theorem transitive_le [Preorder α] : Transitive (@LE.le α _) := diff --git a/Mathlib/SetTheory/Cardinal/Aleph.lean b/Mathlib/SetTheory/Cardinal/Aleph.lean index dafc9ee33dced..ca79284dac29a 100644 --- a/Mathlib/SetTheory/Cardinal/Aleph.lean +++ b/Mathlib/SetTheory/Cardinal/Aleph.lean @@ -541,7 +541,7 @@ def alephIdx.relIso : @RelIso Cardinal.{u} Ordinal.{u} (· < ·) (· < ·) := def alephIdx : Cardinal → Ordinal := aleph'.symm -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephIdx := rfl @@ -555,7 +555,7 @@ theorem alephIdx.relIso_coe : (alephIdx.relIso : Cardinal → Ordinal) = alephId def Aleph'.relIso := aleph' -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'.relIso_coe : (Aleph'.relIso : Ordinal → Cardinal) = aleph' := rfl @@ -571,11 +571,11 @@ theorem aleph'_le {o₁ o₂ : Ordinal} : aleph' o₁ ≤ aleph' o₂ ↔ o₁ theorem aleph'_max (o₁ o₂ : Ordinal) : aleph' (max o₁ o₂) = max (aleph' o₁) (aleph' o₂) := aleph'.monotone.map_max -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem aleph'_alephIdx (c : Cardinal) : aleph' c.alephIdx = c := Cardinal.alephIdx.relIso.toEquiv.symm_apply_apply c -@[deprecated (since := "2024-08-28")] +@[deprecated "No deprecation message was provided." (since := "2024-08-28")] theorem alephIdx_aleph' (o : Ordinal) : (aleph' o).alephIdx = o := Cardinal.alephIdx.relIso.toEquiv.apply_symm_apply o @@ -608,7 +608,7 @@ theorem aleph'_limit {o : Ordinal} (ho : o.IsLimit) : aleph' o = ⨆ a : Iio o, theorem aleph'_omega0 : aleph' ω = ℵ₀ := preAleph_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias aleph'_omega := aleph'_omega0 /-- `aleph'` and `aleph_idx` form an equivalence between `Ordinal` and `Cardinal` -/ @@ -638,7 +638,7 @@ theorem aleph'_isNormal : IsNormal (ord ∘ aleph') := -- They should also use `¬ BddAbove` instead of `Unbounded (· < ·)`. /-- Ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b } := unbounded_lt_iff.2 fun a => ⟨_, @@ -646,16 +646,16 @@ theorem ord_card_unbounded : Unbounded (· < ·) { b : Ordinal | b.card.ord = b dsimp rw [card_ord], (lt_ord_succ_card a).le⟩⟩ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph'_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) : ∃ a, (aleph' a).ord = o := ⟨aleph'.symm o.card, by simpa using ho⟩ /-- Infinite ordinals that are cardinals are unbounded. -/ -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem ord_card_unbounded' : Unbounded (· < ·) { b : Ordinal | b.card.ord = b ∧ ω ≤ b } := (unbounded_lt_inter_le ω).2 ord_card_unbounded -@[deprecated (since := "2024-09-24")] +@[deprecated "No deprecation message was provided." (since := "2024-09-24")] theorem eq_aleph_of_eq_card_ord {o : Ordinal} (ho : o.card.ord = o) (ho' : ω ≤ o) : ∃ a, (ℵ_ a).ord = o := by cases' eq_aleph'_of_eq_card_ord ho with a ha diff --git a/Mathlib/SetTheory/Cardinal/Basic.lean b/Mathlib/SetTheory/Cardinal/Basic.lean index 18f50f95db1cb..e0105b3587892 100644 --- a/Mathlib/SetTheory/Cardinal/Basic.lean +++ b/Mathlib/SetTheory/Cardinal/Basic.lean @@ -145,7 +145,7 @@ protected theorem eq : #α = #β ↔ Nonempty (α ≃ β) := Quotient.eq' /-- Avoid using `Quotient.mk` to construct a `Cardinal` directly -/ -@[deprecated (since := "2024-10-24")] +@[deprecated "No deprecation message was provided." (since := "2024-10-24")] theorem mk'_def (α : Type u) : @Eq Cardinal ⟦α⟧ #α := rfl diff --git a/Mathlib/SetTheory/Cardinal/Cofinality.lean b/Mathlib/SetTheory/Cardinal/Cofinality.lean index 87b3654599332..9a8363e186c9a 100644 --- a/Mathlib/SetTheory/Cardinal/Cofinality.lean +++ b/Mathlib/SetTheory/Cardinal/Cofinality.lean @@ -114,7 +114,7 @@ def StrictOrder.cof (r : α → α → Prop) : Cardinal := Order.cof (swap rᶜ) /-- The set in the definition of `Order.StrictOrder.cof` is nonempty. -/ -@[deprecated (since := "2024-10-22")] +@[deprecated "No deprecation message was provided." (since := "2024-10-22")] theorem StrictOrder.cof_nonempty (r : α → α → Prop) [IsIrrefl α r] : { c | ∃ S : Set α, Unbounded r S ∧ #S = c }.Nonempty := @Order.cof_nonempty α _ (IsRefl.swap rᶜ) diff --git a/Mathlib/SetTheory/Game/Ordinal.lean b/Mathlib/SetTheory/Game/Ordinal.lean index d281b72b36a02..cf5a7e8f4cf97 100644 --- a/Mathlib/SetTheory/Game/Ordinal.lean +++ b/Mathlib/SetTheory/Game/Ordinal.lean @@ -35,7 +35,7 @@ noncomputable def toPGame (o : Ordinal.{u}) : PGame.{u} := termination_by o decreasing_by exact ((enumIsoToType o).symm x).prop -@[deprecated (since := "2024-09-22")] +@[deprecated "No deprecation message was provided." (since := "2024-09-22")] theorem toPGame_def (o : Ordinal) : o.toPGame = ⟨o.toType, PEmpty, fun x => ((enumIsoToType o).symm x).val.toPGame, PEmpty.elim⟩ := by rw [toPGame] diff --git a/Mathlib/SetTheory/Ordinal/Arithmetic.lean b/Mathlib/SetTheory/Ordinal/Arithmetic.lean index 6cc7dfe51df9e..3563908538d09 100644 --- a/Mathlib/SetTheory/Ordinal/Arithmetic.lean +++ b/Mathlib/SetTheory/Ordinal/Arithmetic.lean @@ -605,14 +605,14 @@ theorem one_add_omega0 : 1 + ω = ω := by cases a <;> cases b <;> intro H <;> cases' H with _ _ H _ _ H <;> [exact H.elim; exact Nat.succ_pos _; exact Nat.succ_lt_succ H] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_omega := one_add_omega0 @[simp] theorem one_add_of_omega0_le {o} (h : ω ≤ o) : 1 + o = o := by rw [← Ordinal.add_sub_cancel_of_le h, ← add_assoc, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_add_of_omega_le := one_add_of_omega0_le /-! ### Multiplication of ordinals -/ @@ -1185,7 +1185,7 @@ def sup {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal.{max u v} := iSup f set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sSup_eq_sup {ι : Type u} (f : ι → Ordinal.{max u v}) : sSup (Set.range f) = sup.{_, v} f := rfl @@ -1252,14 +1252,15 @@ protected theorem lt_iSup_iff {ι} {f : ι → Ordinal.{u}} {a : Ordinal.{u}} [S a < iSup f ↔ ∃ i, a < f i := lt_ciSup_iff' (bddAbove_of_small _) -@[deprecated (since := "2024-11-12")] alias lt_iSup := lt_iSup_iff +@[deprecated "No deprecation message was provided." (since := "2024-11-12")] +alias lt_iSup := lt_iSup_iff set_option linter.deprecated false in @[deprecated Ordinal.lt_iSup (since := "2024-08-27")] theorem lt_sup {ι : Type u} {f : ι → Ordinal.{max u v}} {a} : a < sup.{_, v} f ↔ ∃ i, a < f i := by simpa only [not_forall, not_le] using not_congr (@sup_le_iff.{_, v} _ f a) -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem ne_iSup_iff_lt_iSup {ι : Type u} {f : ι → Ordinal.{max u v}} : (∀ i, f i ≠ iSup f) ↔ ∀ i, f i < iSup f := forall_congr' fun i => (Ordinal.le_iSup f i).lt_iff_ne.symm @@ -1377,7 +1378,7 @@ theorem unbounded_range_of_sup_ge {α β : Type u} (r : α → α → Prop) [IsW unbounded_range_of_le_iSup r f h set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem le_sup_shrink_equiv {s : Set Ordinal.{u}} (hs : Small.{u} s) (a) (ha : a ∈ s) : a ≤ sup.{u, u} fun x => ((@equivShrink s hs).symm x).val := by convert le_sup.{u, u} (fun x => ((@equivShrink s hs).symm x).val) ((@equivShrink s hs) ⟨a, ha⟩) @@ -1422,7 +1423,7 @@ theorem IsNormal.apply_of_isLimit {f : Ordinal.{u} → Ordinal.{v}} (H : IsNorma rw [← H.map_iSup, ho.iSup_Iio] set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_eq_sSup {s : Set Ordinal.{u}} (hs : Small.{u} s) : (sup.{u, u} fun x => (@equivShrink s hs).symm x) = sSup s := let hs' := bddAbove_iff_small.2 hs @@ -1986,12 +1987,12 @@ theorem IsNormal.eq_iff_zero_and_succ {f g : Ordinal.{u} → Ordinal.{u}} (hf : Deprecated. If you need this value explicitly, write it in terms of `iSup`. If you just want an upper bound for the image of `op`, use that `Iio a ×ˢ Iio b` is a small set. -/ -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] def blsub₂ (o₁ o₂ : Ordinal) (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) : Ordinal := lsub (fun x : o₁.toType × o₂.toType => op (typein_lt_self x.1) (typein_lt_self x.2)) -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem lt_blsub₂ {o₁ o₂ : Ordinal} (op : {a : Ordinal} → (a < o₁) → {b : Ordinal} → (b < o₂) → Ordinal) {a b : Ordinal} (ha : a < o₁) (hb : b < o₂) : op ha hb < blsub₂ o₁ o₂ op := by @@ -2012,34 +2013,34 @@ set_option linter.deprecated false def mex {ι : Type u} (f : ι → Ordinal.{max u v}) : Ordinal := sInf (Set.range f)ᶜ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_not_mem_range {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ∉ Set.range f := csInf_mem (nonempty_compl_range.{_, v} f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_mex_of_forall {ι : Type u} {f : ι → Ordinal.{max u v}} {a : Ordinal} (H : ∀ b < a, ∃ i, f i = b) : a ≤ mex.{_, v} f := by by_contra! h exact mex_not_mem_range f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_mex {ι : Type u} (f : ι → Ordinal.{max u v}) : ∀ i, f i ≠ mex.{_, v} f := by simpa using mex_not_mem_range.{_, v} f -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_of_ne {ι} {f : ι → Ordinal} {a} (ha : ∀ i, f i ≠ a) : mex f ≤ a := csInf_le' (by simp [ha]) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_mex {ι} {f : ι → Ordinal} {a} (ha : a < mex f) : ∃ i, f i = a := by by_contra! ha' exact ha.not_le (mex_le_of_ne ha') -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_le_lsub {ι : Type u} (f : ι → Ordinal.{max u v}) : mex.{_, v} f ≤ lsub.{_, v} f := csInf_le' (lsub_not_mem_range f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem mex_monotone {α β : Type u} {f : α → Ordinal.{max u v}} {g : β → Ordinal.{max u v}} (h : Set.range f ⊆ Set.range g) : mex.{_, v} f ≤ mex.{_, v} g := by refine mex_le_of_ne fun i hi => ?_ @@ -2072,18 +2073,18 @@ theorem mex_lt_ord_succ_mk {ι : Type u} (f : ι → Ordinal.{u}) : def bmex (o : Ordinal) (f : ∀ a < o, Ordinal) : Ordinal := mex (familyOfBFamily o f) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_not_mem_brange {o : Ordinal} (f : ∀ a < o, Ordinal) : bmex o f ∉ brange o f := by rw [← range_familyOfBFamily] apply mex_not_mem_range -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem le_bmex_of_forall {o : Ordinal} (f : ∀ a < o, Ordinal) {a : Ordinal} (H : ∀ b < a, ∃ i hi, f i hi = b) : a ≤ bmex o f := by by_contra! h exact bmex_not_mem_brange f (H _ h) -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : f i hi ≠ bmex.{_, v} o f := by convert (config := {transparency := .default}) @@ -2091,23 +2092,23 @@ theorem ne_bmex {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) {i} (hi) : -- Porting note: `familyOfBFamily_enum` → `typein_enum` rw [typein_enum] -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_of_ne {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : ∀ i hi, f i hi ≠ a) : bmex o f ≤ a := mex_le_of_ne fun _i => ha _ _ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem exists_of_lt_bmex {o : Ordinal} {f : ∀ a < o, Ordinal} {a} (ha : a < bmex o f) : ∃ i hi, f i hi = a := by cases' exists_of_lt_mex ha with i hi exact ⟨_, typein_lt_self i, hi⟩ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_le_blsub {o : Ordinal.{u}} (f : ∀ a < o, Ordinal.{max u v}) : bmex.{_, v} o f ≤ blsub.{_, v} o f := mex_le_lsub _ -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem bmex_monotone {o o' : Ordinal.{u}} {f : ∀ a < o, Ordinal.{max u v}} {g : ∀ a < o', Ordinal.{max u v}} (h : brange o f ⊆ brange o' g) : bmex.{_, v} o f ≤ bmex.{_, v} o' g := @@ -2165,7 +2166,7 @@ theorem one_add_natCast (m : ℕ) : 1 + (m : Ordinal) = succ m := by rw [← Nat.cast_one, ← Nat.cast_add, add_comm] rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias one_add_nat_cast := one_add_natCast -- See note [no_index around OfNat.ofNat] @@ -2179,43 +2180,43 @@ theorem natCast_mul (m : ℕ) : ∀ n : ℕ, ((m * n : ℕ) : Ordinal) = m * n | 0 => by simp | n + 1 => by rw [Nat.mul_succ, Nat.cast_add, natCast_mul m n, Nat.cast_succ, mul_add_one] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mul := natCast_mul @[deprecated Nat.cast_le (since := "2024-10-17")] theorem natCast_le {m n : ℕ} : (m : Ordinal) ≤ n ↔ m ≤ n := Nat.cast_le -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_le := natCast_le @[deprecated Nat.cast_inj (since := "2024-10-17")] theorem natCast_inj {m n : ℕ} : (m : Ordinal) = n ↔ m = n := Nat.cast_inj -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_inj := natCast_inj @[deprecated Nat.cast_lt (since := "2024-10-17")] theorem natCast_lt {m n : ℕ} : (m : Ordinal) < n ↔ m < n := Nat.cast_lt -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_lt := natCast_lt @[deprecated Nat.cast_eq_zero (since := "2024-10-17")] theorem natCast_eq_zero {n : ℕ} : (n : Ordinal) = 0 ↔ n = 0 := Nat.cast_eq_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_eq_zero := natCast_eq_zero @[deprecated Nat.cast_ne_zero (since := "2024-10-17")] theorem natCast_ne_zero {n : ℕ} : (n : Ordinal) ≠ 0 ↔ n ≠ 0 := Nat.cast_ne_zero -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_ne_zero := natCast_ne_zero @[deprecated Nat.cast_pos' (since := "2024-10-17")] theorem natCast_pos {n : ℕ} : (0 : Ordinal) < n ↔ 0 < n := Nat.cast_pos' -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_pos := natCast_pos @[simp, norm_cast] @@ -2226,7 +2227,7 @@ theorem natCast_sub (m n : ℕ) : ((m - n : ℕ) : Ordinal) = m - n := by · apply (add_left_cancel n).1 rw [← Nat.cast_add, add_tsub_cancel_of_le h, Ordinal.add_sub_cancel_of_le (Nat.cast_le.2 h)] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_sub := natCast_sub @[simp, norm_cast] @@ -2241,7 +2242,7 @@ theorem natCast_div (m n : ℕ) : ((m / n : ℕ) : Ordinal) = m / n := by ← Nat.div_lt_iff_lt_mul (Nat.pos_of_ne_zero hn)] apply Nat.lt_succ_self -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_div := natCast_div @[simp, norm_cast] @@ -2249,7 +2250,7 @@ theorem natCast_mod (m n : ℕ) : ((m % n : ℕ) : Ordinal) = m % n := by rw [← add_left_cancel, div_add_mod, ← natCast_div, ← natCast_mul, ← Nat.cast_add, Nat.div_add_mod] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_mod := natCast_mod @[simp] @@ -2257,7 +2258,7 @@ theorem lift_natCast : ∀ n : ℕ, lift.{u, v} n = n | 0 => by simp | n + 1 => by simp [lift_natCast n] -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias lift_nat_cast := lift_natCast -- See note [no_index around OfNat.ofNat] @@ -2292,13 +2293,13 @@ theorem lt_add_of_limit {a b c : Ordinal.{u}} (h : IsLimit c) : theorem lt_omega0 {o : Ordinal} : o < ω ↔ ∃ n : ℕ, o = n := by simp_rw [← Cardinal.ord_aleph0, Cardinal.lt_ord, lt_aleph0, card_eq_nat] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias lt_omega := lt_omega0 theorem nat_lt_omega0 (n : ℕ) : ↑n < ω := lt_omega0.2 ⟨_, rfl⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nat_lt_omega := nat_lt_omega0 theorem omega0_pos : 0 < ω := @@ -2307,12 +2308,12 @@ theorem omega0_pos : 0 < ω := theorem omega0_ne_zero : ω ≠ 0 := omega0_pos.ne' -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_ne_zero := omega0_ne_zero theorem one_lt_omega0 : 1 < ω := by simpa only [Nat.cast_one] using nat_lt_omega0 1 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias one_lt_omega := one_lt_omega0 theorem isLimit_omega0 : IsLimit ω := @@ -2320,10 +2321,10 @@ theorem isLimit_omega0 : IsLimit ω := let ⟨n, e⟩ := lt_omega0.1 h rw [e]; exact nat_lt_omega0 (n + 1)⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias omega0_isLimit := isLimit_omega0 -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_isLimit := isLimit_omega0 theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := @@ -2332,7 +2333,7 @@ theorem omega0_le {o : Ordinal} : ω ≤ o ↔ ∀ n : ℕ, ↑n ≤ o := let ⟨n, e⟩ := lt_omega0.1 h rw [e, ← succ_le_iff]; exact H (n + 1)⟩ -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le := omega0_le @[simp] @@ -2344,7 +2345,7 @@ set_option linter.deprecated false in theorem sup_natCast : sup Nat.cast = ω := iSup_natCast -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias sup_nat_cast := sup_natCast theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o @@ -2354,7 +2355,7 @@ theorem nat_lt_limit {o} (h : IsLimit o) : ∀ n : ℕ, ↑n < o theorem omega0_le_of_isLimit {o} (h : IsLimit o) : ω ≤ o := omega0_le.2 fun n => le_of_lt <| nat_lt_limit h n -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias omega_le_of_isLimit := omega0_le_of_isLimit theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ a := by @@ -2372,7 +2373,7 @@ theorem isLimit_iff_omega0_dvd {a : Ordinal} : IsLimit a ↔ a ≠ 0 ∧ ω ∣ intro e simp only [e, mul_zero] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias isLimit_iff_omega_dvd := isLimit_iff_omega0_dvd theorem add_mul_limit_aux {a b c : Ordinal} (ba : b + a = a) (l : IsLimit c) @@ -2416,7 +2417,7 @@ theorem add_le_of_forall_add_lt {a b c : Ordinal} (hb : 0 < b) (h : ∀ d < b, a theorem IsNormal.apply_omega0 {f : Ordinal.{u} → Ordinal.{v}} (hf : IsNormal f) : ⨆ n : ℕ, f n = f ω := by rw [← iSup_natCast, hf.map_iSup] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias IsNormal.apply_omega := IsNormal.apply_omega0 @[simp] @@ -2460,7 +2461,7 @@ theorem isLimit_ord {c} (co : ℵ₀ ≤ c) : (ord c).IsLimit := by · rw [ord_aleph0] exact Ordinal.isLimit_omega0 -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] alias ord_isLimit := isLimit_ord theorem noMaxOrder {c} (h : ℵ₀ ≤ c) : NoMaxOrder c.ord.toType := diff --git a/Mathlib/SetTheory/Ordinal/Basic.lean b/Mathlib/SetTheory/Ordinal/Basic.lean index 5a31bfe919ee9..ec9a34419fbf4 100644 --- a/Mathlib/SetTheory/Ordinal/Basic.lean +++ b/Mathlib/SetTheory/Ordinal/Basic.lean @@ -387,7 +387,7 @@ def typein (r : α → α → Prop) [IsWellOrder α r] : @PrincipalSeg α Ordina alias typein.principalSeg := typein set_option linter.deprecated false in -@[deprecated (since := "2024-10-09")] +@[deprecated "No deprecation message was provided." (since := "2024-10-09")] theorem typein.principalSeg_coe (r : α → α → Prop) [IsWellOrder α r] : (typein.principalSeg r : α → Ordinal) = typein r := rfl @@ -513,7 +513,7 @@ noncomputable def enumIsoToType (o : Ordinal) : Set.Iio o ≃o o.toType where right_inv _ := enum_typein _ _ map_rel_iff' := enum_le_enum' _ -@[deprecated (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias enumIsoOut := enumIsoToType instance small_Iio (o : Ordinal.{u}) : Small.{u} (Iio o) := @@ -925,7 +925,7 @@ theorem card_succ (o : Ordinal) : card (succ o) = card o + 1 := by theorem natCast_succ (n : ℕ) : ↑n.succ = succ (n : Ordinal) := rfl -@[deprecated (since := "2024-04-17")] +@[deprecated "No deprecation message was provided." (since := "2024-04-17")] alias nat_cast_succ := natCast_succ instance uniqueIioOne : Unique (Iio (1 : Ordinal)) where @@ -1213,7 +1213,7 @@ alias card_typein_out_lt := card_typein_toType_lt theorem mk_Iio_ord_toType {c : Cardinal} (i : c.ord.toType) : #(Iio i) < c := card_typein_toType_lt c i -@[deprecated (since := "2024-08-26")] +@[deprecated "No deprecation message was provided." (since := "2024-08-26")] alias mk_Iio_ord_out_α := mk_Iio_ord_toType theorem ord_injective : Injective ord := by diff --git a/Mathlib/SetTheory/Ordinal/Enum.lean b/Mathlib/SetTheory/Ordinal/Enum.lean index 4e5d4d38a7cf6..e8b3bb9314bc5 100644 --- a/Mathlib/SetTheory/Ordinal/Enum.lean +++ b/Mathlib/SetTheory/Ordinal/Enum.lean @@ -33,7 +33,7 @@ termination_by o variable {s : Set Ordinal.{u}} -@[deprecated (since := "2024-09-20")] +@[deprecated "No deprecation message was provided." (since := "2024-09-20")] theorem enumOrd_def (o : Ordinal.{u}) : enumOrd s o = sInf (s ∩ { b | ∀ c, c < o → enumOrd s c < b }) := by rw [enumOrd] diff --git a/Mathlib/SetTheory/Ordinal/FixedPoint.lean b/Mathlib/SetTheory/Ordinal/FixedPoint.lean index 36b50a052b83d..84f2f8db36b0d 100644 --- a/Mathlib/SetTheory/Ordinal/FixedPoint.lean +++ b/Mathlib/SetTheory/Ordinal/FixedPoint.lean @@ -52,7 +52,7 @@ least `a`, and `Ordinal.nfpFamily_le_fp` shows this is the least ordinal with th def nfpFamily (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : Ordinal := ⨆ i, List.foldr f a i -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpFamily_eq_sup (f : ι → Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : nfpFamily f a = ⨆ i, List.foldr f a i := rfl @@ -238,48 +238,48 @@ def nfpBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal.{ Ordinal.{max u v} → Ordinal.{max u v} := nfpFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_nfpFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : nfpBFamily.{u, v} o f = nfpFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem foldr_le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a l) : List.foldr (familyOfBFamily o f) a l ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_nfpBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) (a) : a ≤ nfpBFamily.{u, v} o f a := Ordinal.le_iSup (fun _ ↦ List.foldr _ a _) [] -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem lt_nfpBFamily {a b} : a < nfpBFamily.{u, v} o f b ↔ ∃ l, a < List.foldr (familyOfBFamily o f) b l := Ordinal.lt_iSup_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_iff {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : nfpBFamily.{u, v} o f a ≤ b ↔ ∀ l, List.foldr (familyOfBFamily o f) a l ≤ b := Ordinal.iSup_le_iff -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le {o : Ordinal} {f : ∀ b < o, Ordinal → Ordinal} {a b} : (∀ l, List.foldr (familyOfBFamily o f) a l ≤ b) → nfpBFamily.{u, v} o f a ≤ b := Ordinal.iSup_le -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_monotone (hf : ∀ i hi, Monotone (f i hi)) : Monotone (nfpBFamily.{u, v} o f) := nfpFamily_monotone fun _ => hf _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily (H : ∀ i hi, IsNormal (f i hi)) {a b} (hb : b < nfpBFamily.{u, v} o f a) (i hi) : f i hi b < nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] apply apply_lt_nfpFamily (fun _ => H _ _) hb -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b < nfpBFamily.{u, v} o f a) ↔ b < nfpBFamily.{u, v} o f a := ⟨fun h => by @@ -287,19 +287,19 @@ theorem apply_lt_nfpBFamily_iff (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) refine (apply_lt_nfpFamily_iff ?_).1 fun _ => h _ _ exact fun _ => H _ _, apply_lt_nfpBFamily H⟩ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_apply (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∃ i hi, nfpBFamily.{u, v} o f a ≤ f i hi b) ↔ nfpBFamily.{u, v} o f a ≤ b := by rw [← not_iff_not] push_neg exact apply_lt_nfpBFamily_iff.{u, v} ho H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_le_fp (H : ∀ i hi, Monotone (f i hi)) {a b} (ab : a ≤ b) (h : ∀ i hi, f i hi b ≤ b) : nfpBFamily.{u, v} o f a ≤ b := nfpFamily_le_fp (fun _ => H _ _) ab fun _ => h _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : f i hi (nfpBFamily.{u, v} o f a) = nfpBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -307,7 +307,7 @@ theorem nfpBFamily_fp {i hi} (H : IsNormal (f i hi)) (a) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a b} : (∀ i hi, f i hi b ≤ nfpBFamily.{u, v} o f a) ↔ b ≤ nfpBFamily.{u, v} o f a := by refine ⟨fun h => ?_, fun h i hi => ?_⟩ @@ -316,13 +316,13 @@ theorem apply_le_nfpBFamily (ho : o ≠ 0) (H : ∀ i hi, IsNormal (f i hi)) {a · rw [← nfpBFamily_fp (H i hi)] exact (H i hi).monotone h -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem nfpBFamily_eq_self {a} (h : ∀ i hi, f i hi a = a) : nfpBFamily.{u, v} o f a = a := nfpFamily_eq_self fun _ => h _ _ /-- A generalization of the fixed point lemma for normal functions: any family of normal functions has an unbounded set of common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem not_bddAbove_fp_bfamily (H : ∀ i hi, IsNormal (f i hi)) : ¬ BddAbove (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [not_bddAbove_iff] @@ -347,12 +347,12 @@ def derivBFamily (o : Ordinal.{u}) (f : ∀ b < o, Ordinal.{max u v} → Ordinal Ordinal.{max u v} → Ordinal.{max u v} := derivFamily (familyOfBFamily o f) -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_derivFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : derivBFamily.{u, v} o f = derivFamily (familyOfBFamily o f) := rfl -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) : IsNormal (derivBFamily o f) := isNormal_derivFamily _ @@ -360,7 +360,7 @@ theorem isNormal_derivBFamily {o : Ordinal} (f : ∀ b < o, Ordinal → Ordinal) @[deprecated isNormal_derivBFamily (since := "2024-10-11")] alias derivBFamily_isNormal := isNormal_derivBFamily -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : f i hi (derivBFamily.{u, v} o f a) = derivBFamily.{u, v} o f a := by rw [← familyOfBFamily_enum o f] @@ -368,7 +368,7 @@ theorem derivBFamily_fp {i hi} (H : IsNormal (f i hi)) (a : Ordinal) : rw [familyOfBFamily_enum] exact H -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a ≤ a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by unfold derivBFamily @@ -378,7 +378,7 @@ theorem le_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : apply h · exact fun _ => H _ _ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : (∀ i hi, f i hi a = a) ↔ ∃ b, derivBFamily.{u, v} o f b = a := by rw [← le_iff_derivBFamily H] @@ -387,7 +387,7 @@ theorem fp_iff_derivBFamily (H : ∀ i hi, IsNormal (f i hi)) {a} : exact h i hi /-- For a family of normal functions, `Ordinal.derivBFamily` enumerates the common fixed points. -/ -@[deprecated (since := "2024-10-14")] +@[deprecated "No deprecation message was provided." (since := "2024-10-14")] theorem derivBFamily_eq_enumOrd (H : ∀ i hi, IsNormal (f i hi)) : derivBFamily.{u, v} o f = enumOrd (⋂ (i) (hi), Function.fixedPoints (f i hi)) := by rw [eq_comm, eq_enumOrd _ (not_bddAbove_fp_bfamily H)] @@ -428,7 +428,7 @@ theorem iSup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) exact Ordinal.le_iSup _ _ set_option linter.deprecated false in -@[deprecated (since := "2024-08-27")] +@[deprecated "No deprecation message was provided." (since := "2024-08-27")] theorem sup_iterate_eq_nfp (f : Ordinal.{u} → Ordinal.{u}) (a : Ordinal.{u}) : (sup fun n : ℕ => f^[n] a) = nfp f a := by refine le_antisymm ?_ (sup_le fun l => ?_) @@ -580,7 +580,7 @@ theorem nfp_add_eq_mul_omega0 {a b} (hba : b ≤ a * ω) : nfp (a + ·) b = a * exact nfp_monotone (isNormal_add_right a).monotone (Ordinal.zero_le b) · dsimp; rw [← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_add_eq_mul_omega := nfp_add_eq_mul_omega0 theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω ≤ b := by @@ -593,14 +593,14 @@ theorem add_eq_right_iff_mul_omega0_le {a b : Ordinal} : a + b = b ↔ a * ω nth_rw 1 [← this] rwa [← add_assoc, ← mul_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias add_eq_right_iff_mul_omega_le := add_eq_right_iff_mul_omega0_le theorem add_le_right_iff_mul_omega0_le {a b : Ordinal} : a + b ≤ b ↔ a * ω ≤ b := by rw [← add_eq_right_iff_mul_omega0_le] exact (isNormal_add_right a).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias add_le_right_iff_mul_omega_le := add_le_right_iff_mul_omega0_le theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * ω + b := by @@ -612,7 +612,7 @@ theorem deriv_add_eq_mul_omega0_add (a b : Ordinal.{u}) : deriv (a + ·) b = a * · rw [deriv_succ, h, add_succ] exact nfp_eq_self (add_eq_right_iff_mul_omega0_le.2 ((le_add_right _ _).trans (le_succ _))) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias deriv_add_eq_mul_omega_add := deriv_add_eq_mul_omega0_add /-! ### Fixed points of multiplication -/ @@ -645,7 +645,7 @@ theorem nfp_mul_eq_opow_omega0 {a b : Ordinal} (hb : 0 < b) (hba : b ≤ a ^ ω) rw [← nfp_mul_one ha] exact nfp_monotone (isNormal_mul_right ha).monotone (one_le_iff_pos.2 hb) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_mul_eq_opow_omega := nfp_mul_eq_opow_omega0 theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = b) : @@ -659,7 +659,7 @@ theorem eq_zero_or_opow_omega0_le_of_mul_eq_right {a b : Ordinal} (hab : a * b = rw [← Ne, ← one_le_iff_ne_zero] at hb exact nfp_le_fp (isNormal_mul_right ha).monotone hb (le_of_eq hab) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias eq_zero_or_opow_omega_le_of_mul_eq_right := eq_zero_or_opow_omega0_le_of_mul_eq_right theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω ∣ b := by @@ -677,7 +677,7 @@ theorem mul_eq_right_iff_opow_omega0_dvd {a b : Ordinal} : a * b = b ↔ a ^ ω cases' h with c hc rw [hc, ← mul_assoc, ← opow_one_add, one_add_omega0] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias mul_eq_right_iff_opow_omega_dvd := mul_eq_right_iff_opow_omega0_dvd theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : @@ -685,7 +685,7 @@ theorem mul_le_right_iff_opow_omega0_dvd {a b : Ordinal} (ha : 0 < a) : rw [← mul_eq_right_iff_opow_omega0_dvd] exact (isNormal_mul_right ha).le_iff_eq -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias mul_le_right_iff_opow_omega_dvd := mul_le_right_iff_opow_omega0_dvd theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) @@ -705,7 +705,7 @@ theorem nfp_mul_opow_omega0_add {a c : Ordinal} (b) (ha : 0 < a) (hc : 0 < c) rw [add_zero, mul_lt_mul_iff_left (opow_pos ω ha)] at this rwa [succ_le_iff] -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias nfp_mul_opow_omega_add := nfp_mul_opow_omega0_add theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : @@ -718,7 +718,7 @@ theorem deriv_mul_eq_opow_omega0_mul {a : Ordinal.{u}} (ha : 0 < a) (b) : · rw [deriv_succ, h] exact nfp_mul_opow_omega0_add c ha zero_lt_one (one_le_iff_pos.2 (opow_pos _ ha)) -@[deprecated (since := "2024-09-30")] +@[deprecated "No deprecation message was provided." (since := "2024-09-30")] alias deriv_mul_eq_opow_omega_mul := deriv_mul_eq_opow_omega0_mul end Ordinal diff --git a/Mathlib/SetTheory/Ordinal/Principal.lean b/Mathlib/SetTheory/Ordinal/Principal.lean index 5d6417f174042..194abeabc7c50 100644 --- a/Mathlib/SetTheory/Ordinal/Principal.lean +++ b/Mathlib/SetTheory/Ordinal/Principal.lean @@ -129,7 +129,7 @@ theorem not_bddAbove_principal (op : Ordinal → Ordinal → Ordinal) : exact ((le_nfp _ _).trans (ha (principal_nfp_iSup op (succ a)))).not_lt (lt_succ a) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordinal) : Principal op (nfp (fun o' => blsub₂.{u, u, u} o' o' (@fun a _ b _ => op a b)) o) := by intro a b ha hb @@ -147,7 +147,7 @@ theorem principal_nfp_blsub₂ (op : Ordinal → Ordinal → Ordinal) (o : Ordin exact lt_blsub₂ (@fun a _ b _ => op a b) hm (hn.trans_le h) set_option linter.deprecated false in -@[deprecated (since := "2024-10-11")] +@[deprecated "No deprecation message was provided." (since := "2024-10-11")] theorem unbounded_principal (op : Ordinal → Ordinal → Ordinal) : Set.Unbounded (· < ·) { o | Principal op o } := fun o => ⟨_, principal_nfp_blsub₂ op o, (le_nfp _ o).not_lt⟩ diff --git a/Mathlib/SetTheory/ZFC/Basic.lean b/Mathlib/SetTheory/ZFC/Basic.lean index c7a0c35cdfcef..a2b3354096bf0 100644 --- a/Mathlib/SetTheory/ZFC/Basic.lean +++ b/Mathlib/SetTheory/ZFC/Basic.lean @@ -447,12 +447,12 @@ set_option linter.deprecated false /-- Function equivalence is defined so that `f ~ g` iff `∀ x y, x ~ y → f x ~ g y`. This extends to equivalence of `n`-ary functions. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Arity.Equiv : ∀ {n}, OfArity PSet.{u} PSet.{u} n → OfArity PSet.{u} PSet.{u} n → Prop | 0, a, b => PSet.Equiv a b | _ + 1, a, b => ∀ x y : PSet, PSet.Equiv x y → Arity.Equiv (a x) (b y) -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Arity.equiv_const {a : PSet.{u}} : ∀ n, Arity.Equiv (OfArity.const PSet.{u} a n) (OfArity.const PSet.{u} a n) | 0 => Equiv.rfl @@ -460,46 +460,46 @@ theorem Arity.equiv_const {a : PSet.{u}} : /-- `resp n` is the collection of n-ary functions on `PSet` that respect equivalence, i.e. when the inputs are equivalent the output is as well. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp (n) := { x : OfArity PSet.{u} PSet.{u} n // Arity.Equiv x x } -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.inhabited {n} : Inhabited (Resp n) := ⟨⟨OfArity.const _ default _, Arity.equiv_const _⟩⟩ /-- The `n`-ary image of a `(n + 1)`-ary function respecting equivalence as a function respecting equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.f {n} (f : Resp (n + 1)) (x : PSet) : Resp n := ⟨f.1 x, f.2 _ _ <| Equiv.refl x⟩ /-- Function equivalence for functions respecting equivalence. See `PSet.Arity.Equiv`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Resp.Equiv {n} (a b : Resp n) : Prop := Arity.Equiv a.1 b.1 -@[deprecated (since := "2024-09-02"), refl] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), refl] protected theorem Resp.Equiv.refl {n} (a : Resp n) : Resp.Equiv a a := a.2 -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] protected theorem Resp.Equiv.euc : ∀ {n} {a b c : Resp n}, Resp.Equiv a b → Resp.Equiv c b → Resp.Equiv a c | 0, _, _, _, hab, hcb => PSet.Equiv.euc hab hcb | n + 1, a, b, c, hab, hcb => fun x y h => @Resp.Equiv.euc n (a.f x) (b.f y) (c.f y) (hab _ _ h) (hcb _ _ <| PSet.Equiv.refl y) -@[deprecated (since := "2024-09-02"), symm] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), symm] protected theorem Resp.Equiv.symm {n} {a b : Resp n} : Resp.Equiv a b → Resp.Equiv b a := (Resp.Equiv.refl b).euc -@[deprecated (since := "2024-09-02"), trans] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), trans] protected theorem Resp.Equiv.trans {n} {x y z : Resp n} (h1 : Resp.Equiv x y) (h2 : Resp.Equiv y z) : Resp.Equiv x z := h1.euc h2.symm -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] instance Resp.setoid {n} : Setoid (Resp n) := ⟨Resp.Equiv, Resp.Equiv.refl, Resp.Equiv.symm, Resp.Equiv.trans⟩ @@ -611,7 +611,7 @@ set_option linter.deprecated false namespace Resp /-- Helper function for `PSet.eval`. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def evalAux : ∀ {n}, { f : Resp n → OfArity ZFSet.{u} ZFSet.{u} n // ∀ a b : Resp n, Resp.Equiv a b → f a = f b } @@ -626,11 +626,11 @@ def evalAux : evalAux.2 (Resp.f b z) (Resp.f c z) (h _ _ (PSet.Equiv.refl z))⟩ /-- An equivalence-respecting function yields an n-ary ZFC set function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def eval (n) : Resp n → OfArity ZFSet.{u} ZFSet.{u} n := evalAux.1 -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem eval_val {n f x} : (@eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) ⟦x⟧ = eval n (Resp.f f x) := rfl @@ -640,24 +640,25 @@ end Resp /-- A set function is "definable" if it is the image of some n-ary pre-set function. This isn't exactly definability, but is useful as a sufficient condition for functions that have a computable image. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] class inductive Definable (n) : OfArity ZFSet.{u} ZFSet.{u} n → Type (u + 1) | mk (f) : Definable n (Resp.eval n f) -attribute [deprecated (since := "2024-09-02"), instance] Definable.mk +attribute [deprecated "No deprecation message was provided." (since := "2024-09-02"), instance] + Definable.mk /-- The evaluation of a function respecting equivalence is definable, by that same function. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.EqMk {n} (f) : ∀ {s : OfArity ZFSet.{u} ZFSet.{u} n} (_ : Resp.eval _ f = s), Definable n s | _, rfl => ⟨f⟩ /-- Turns a definable function into a function that respects equivalence. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] def Definable.Resp {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [Definable n s], Resp n | _, ⟨f⟩ => f -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] theorem Definable.eq {n} : ∀ (s : OfArity ZFSet.{u} ZFSet.{u} n) [H : Definable n s], (@Definable.Resp n s H).eval _ = s | _, ⟨_⟩ => rfl @@ -670,7 +671,7 @@ open PSet ZFSet set_option linter.deprecated false in /-- All functions are classically definable. -/ -@[deprecated (since := "2024-09-02")] +@[deprecated "No deprecation message was provided." (since := "2024-09-02")] noncomputable def allDefinable : ∀ {n} (F : OfArity ZFSet ZFSet n), Definable n F | 0, F => let p := @Quotient.exists_rep PSet _ F @@ -706,7 +707,7 @@ theorem exact {x y : PSet} : mk x = mk y → PSet.Equiv x y := Quotient.exact set_option linter.deprecated false in -@[deprecated (since := "2024-09-02"), simp] +@[deprecated "No deprecation message was provided." (since := "2024-09-02"), simp] theorem eval_mk {n f x} : (@Resp.eval (n + 1) f : ZFSet → OfArity ZFSet ZFSet n) (mk x) = Resp.eval n (Resp.f f x) := rfl diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 1da40dadc219d..44b6433923f59 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -266,7 +266,7 @@ theorem smul_mem_nhds_smul_iff {t : Set α} (g : G) {a : α} : g • t ∈ 𝓝 @[to_additive] alias ⟨_, smul_mem_nhds_smul⟩ := smul_mem_nhds_smul_iff -@[to_additive (attr := deprecated (since := "2024-08-06"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-06"))] alias smul_mem_nhds := smul_mem_nhds_smul @[to_additive (attr := simp)] diff --git a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean index 91b1a4ca3bcd4..f7373a1da6da0 100644 --- a/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean +++ b/Mathlib/Topology/Algebra/ContinuousMonoidHom.lean @@ -65,7 +65,8 @@ homomorphisms. Deprecated and changed from a `class` to a `structure`. Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead. -/ -@[to_additive (attr := deprecated (since := "2024-10-08"))] +@[to_additive (attr := deprecated "Use `[MonoidHomClass F A B] [ContinuousMapClass F A B]` instead." + (since := "2024-10-08"))] structure ContinuousMonoidHomClass (A B : outParam Type*) [Monoid A] [Monoid B] [TopologicalSpace A] [TopologicalSpace B] [FunLike F A B] extends MonoidHomClass F A B, ContinuousMapClass F A B : Prop diff --git a/Mathlib/Topology/Algebra/Group/Quotient.lean b/Mathlib/Topology/Algebra/Group/Quotient.lean index 4a2b490786409..e46af523bf2ca 100644 --- a/Mathlib/Topology/Algebra/Group/Quotient.lean +++ b/Mathlib/Topology/Algebra/Group/Quotient.lean @@ -76,7 +76,7 @@ instance instLocallyCompactSpace [LocallyCompactSpace G] (N : Subgroup G) : LocallyCompactSpace (G ⧸ N) := QuotientGroup.isOpenQuotientMap_mk.locallyCompactSpace -@[to_additive (attr := deprecated (since := "2024-10-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-10-05"))] theorem continuous_smul₁ (x : G ⧸ N) : Continuous fun g : G => g • x := continuous_id.smul continuous_const @@ -101,7 +101,7 @@ instance instSecondCountableTopology [SecondCountableTopology G] : SecondCountableTopology (G ⧸ N) := ContinuousConstSMul.secondCountableTopology -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem nhds_one_isCountablyGenerated [FirstCountableTopology G] [N.Normal] : (𝓝 (1 : G ⧸ N)).IsCountablyGenerated := inferInstance @@ -117,7 +117,7 @@ instance instTopologicalGroup [N.Normal] : TopologicalGroup (G ⧸ N) where exact continuous_mk.comp continuous_mul continuous_inv := continuous_inv.quotient_map' _ -@[to_additive (attr := deprecated (since := "2024-08-05"))] +@[to_additive (attr := deprecated "No deprecation message was provided." (since := "2024-08-05"))] theorem _root_.topologicalGroup_quotient [N.Normal] : TopologicalGroup (G ⧸ N) := instTopologicalGroup N diff --git a/Mathlib/Topology/LocalAtTarget.lean b/Mathlib/Topology/LocalAtTarget.lean index 9ab72631191fb..af16705f2ce51 100644 --- a/Mathlib/Topology/LocalAtTarget.lean +++ b/Mathlib/Topology/LocalAtTarget.lean @@ -82,7 +82,7 @@ theorem IsClosedMap.restrictPreimage (H : IsClosedMap f) (s : Set β) : simpa [isClosed_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isClosedMap (s : Set β) (H : IsClosedMap f) : IsClosedMap (s.restrictPreimage f) := H.restrictPreimage s @@ -94,7 +94,7 @@ theorem IsOpenMap.restrictPreimage (H : IsOpenMap f) (s : Set β) : simpa [isOpen_induced_iff] exact fun u hu e => ⟨f '' u, H u hu, by simp [← e, image_restrictPreimage]⟩ -@[deprecated (since := "2024-04-02")] +@[deprecated "No deprecation message was provided." (since := "2024-04-02")] theorem Set.restrictPreimage_isOpenMap (s : Set β) (H : IsOpenMap f) : IsOpenMap (s.restrictPreimage f) := H.restrictPreimage s diff --git a/Mathlib/Topology/MetricSpace/Polish.lean b/Mathlib/Topology/MetricSpace/Polish.lean index d1e851e3c7d44..93c053adc16d7 100644 --- a/Mathlib/Topology/MetricSpace/Polish.lean +++ b/Mathlib/Topology/MetricSpace/Polish.lean @@ -102,7 +102,7 @@ instance (priority := 100) instMetrizableSpace (α : Type*) [TopologicalSpace α letI := upgradePolishSpace α infer_instance -@[deprecated (since := "2024-02-23")] +@[deprecated "No deprecation message was provided." (since := "2024-02-23")] theorem t2Space (α : Type*) [TopologicalSpace α] [PolishSpace α] : T2Space α := inferInstance /-- A countable product of Polish spaces is Polish. -/ diff --git a/Mathlib/Topology/NoetherianSpace.lean b/Mathlib/Topology/NoetherianSpace.lean index 64689fd04f3f0..931e8748cbfce 100644 --- a/Mathlib/Topology/NoetherianSpace.lean +++ b/Mathlib/Topology/NoetherianSpace.lean @@ -98,7 +98,7 @@ theorem noetherianSpace_iff_isCompact : NoetherianSpace α ↔ ∀ s : Set α, I instance [NoetherianSpace α] : WellFoundedLT (Closeds α) := Iff.mp ((noetherianSpace_TFAE α).out 0 1) ‹_› -@[deprecated (since := "2024-10-07")] +@[deprecated "No deprecation message was provided." (since := "2024-10-07")] theorem NoetherianSpace.wellFounded_closeds [NoetherianSpace α] : WellFounded fun s t : Closeds α => s < t := wellFounded_lt diff --git a/Mathlib/Topology/Order/OrderClosed.lean b/Mathlib/Topology/Order/OrderClosed.lean index 7ded140f87a07..46e8626f54295 100644 --- a/Mathlib/Topology/Order/OrderClosed.lean +++ b/Mathlib/Topology/Order/OrderClosed.lean @@ -368,7 +368,7 @@ variable [TopologicalSpace α] [Preorder α] [ClosedIciTopology α] {f : β → theorem isClosed_Ici {a : α} : IsClosed (Ici a) := ClosedIciTopology.isClosed_Ici a -@[deprecated (since := "2024-02-15")] +@[deprecated "No deprecation message was provided." (since := "2024-02-15")] lemma ClosedIciTopology.isClosed_ge' (a : α) : IsClosed {x | a ≤ x} := isClosed_Ici a export ClosedIciTopology (isClosed_ge') From 7113817ab894ff56fb3a99f2efdff55f8da6df16 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 25 Nov 2024 13:34:39 +0000 Subject: [PATCH 145/250] chore: change the definition of `List.finRange` (#19447) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit François Dorais has been [working](https://github.com/leanprover-community/batteries/pull/1055) on upstreaming `List.finRange`, but wants to change the definition at the same time. That was running into difficulties, which hopefully are resolved here. --- Archive/Imo/Imo2024Q5.lean | 2 +- .../Complex/UpperHalfPlane/Basic.lean | 12 ++-- Mathlib/Analysis/Convex/StoneSeparation.lean | 4 +- .../Analysis/InnerProductSpace/TwoDim.lean | 33 +++++----- Mathlib/Data/List/Defs.lean | 12 +++- Mathlib/Data/List/FinRange.lean | 64 +++++++++++++++++-- Mathlib/Data/List/Range.lean | 49 -------------- Mathlib/Data/List/Sublists.lean | 3 +- .../AffineSpace/Independent.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 18 +++--- .../NumberField/CanonicalEmbedding/Basic.lean | 20 +++--- Mathlib/Order/RelSeries.lean | 4 +- Mathlib/RingTheory/Complex.lean | 8 +-- Mathlib/SetTheory/Cardinal/Subfield.lean | 2 +- 14 files changed, 127 insertions(+), 106 deletions(-) diff --git a/Archive/Imo/Imo2024Q5.lean b/Archive/Imo/Imo2024Q5.lean index 7b090f48cc987..5cbdf23de02f8 100644 --- a/Archive/Imo/Imo2024Q5.lean +++ b/Archive/Imo/Imo2024Q5.lean @@ -509,7 +509,7 @@ lemma Strategy.play_two (s : Strategy N) (m : MonsterData N) {k : ℕ} (hk : 2 < fin_cases i · rfl · have h : (1 : Fin 2) = Fin.last 1 := rfl - simp only [Fin.snoc_zero, Nat.reduceAdd, Fin.mk_one, Fin.isValue, Matrix.cons_val_one, + simp only [Fin.snoc_zero, Nat.reduceAdd, Fin.mk_one, Fin.isValue, id_eq, Matrix.cons_val_one, Matrix.head_cons] simp only [h, Fin.snoc_last] convert rfl diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 0ac5af563ff8a..6a3125c87d0c9 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -243,9 +243,9 @@ theorem denom_cocycle (x y : GL(2, ℝ)⁺) (z : ℍ) : denom (x * y) z = denom x (smulAux y z) * denom y z := by change _ = (_ * (_ / _) + _) * _ field_simp [denom_ne_zero] - simp only [Matrix.mul_apply, dotProduct, Fin.sum_univ_succ, denom, num, Subgroup.coe_mul, - GeneralLinearGroup.coe_mul, Fintype.univ_ofSubsingleton, Fin.mk_zero, Finset.sum_singleton, - Fin.succ_zero_eq_one, Complex.ofReal_add, Complex.ofReal_mul] + simp only [denom, Subgroup.coe_mul, Fin.isValue, Units.val_mul, mul_apply, Fin.sum_univ_succ, + Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, Fin.succ_zero_eq_one, + Complex.ofReal_add, Complex.ofReal_mul, num] ring theorem mul_smul' (x y : GL(2, ℝ)⁺) (z : ℍ) : smulAux (x * y) z = smulAux x (smulAux y z) := by @@ -254,9 +254,9 @@ theorem mul_smul' (x y : GL(2, ℝ)⁺) (z : ℍ) : smulAux (x * y) z = smulAux change _ / _ = (_ * (_ / _) + _) / _ rw [denom_cocycle] field_simp [denom_ne_zero] - simp only [Matrix.mul_apply, dotProduct, Fin.sum_univ_succ, num, denom, Subgroup.coe_mul, - GeneralLinearGroup.coe_mul, Fintype.univ_ofSubsingleton, Fin.mk_zero, Finset.sum_singleton, - Fin.succ_zero_eq_one, Complex.ofReal_add, Complex.ofReal_mul] + simp only [num, Subgroup.coe_mul, Fin.isValue, Units.val_mul, mul_apply, Fin.sum_univ_succ, + Finset.univ_unique, Fin.default_eq_zero, Finset.sum_singleton, Fin.succ_zero_eq_one, + Complex.ofReal_add, Complex.ofReal_mul, denom] ring /-- The action of `GLPos 2 ℝ` on the upper half-plane by fractional linear transformations. -/ diff --git a/Mathlib/Analysis/Convex/StoneSeparation.lean b/Mathlib/Analysis/Convex/StoneSeparation.lean index 76a95705677f6..5a52d98fda2ea 100644 --- a/Mathlib/Analysis/Convex/StoneSeparation.lean +++ b/Mathlib/Analysis/Convex/StoneSeparation.lean @@ -67,8 +67,8 @@ theorem not_disjoint_segment_convexHull_triple {p q u v x y z : E} (hz : z ∈ s ((az * av * bu) • p + ((bz * au * bv) • q + (au * av) • (az • x + bz • y))) · module congr 3 - simp only [w, z, smul_add, List.foldr, Matrix.cons_val_succ', Fin.mk_one, - Matrix.cons_val_one, Matrix.head_cons, add_zero] + simp only [smul_add, List.foldr, Fin.reduceFinMk, id_eq, Fin.isValue, Matrix.cons_val_two, + Nat.succ_eq_add_one, Nat.reduceAdd, Matrix.tail_cons, Matrix.head_cons, add_zero, w, z] /-- **Stone's Separation Theorem** -/ theorem exists_convex_convex_compl_subset (hs : Convex 𝕜 s) (ht : Convex 𝕜 t) (hst : Disjoint s t) : diff --git a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean index b07b1d09381f7..afdbc274d603b 100644 --- a/Mathlib/Analysis/InnerProductSpace/TwoDim.lean +++ b/Mathlib/Analysis/InnerProductSpace/TwoDim.lean @@ -354,14 +354,16 @@ theorem inner_mul_inner_add_areaForm_mul_areaForm' (a x : E) : apply (o.basisRightAngleRotation a ha).ext intro i fin_cases i - · simp only [Fin.mk_zero, coe_basisRightAngleRotation, Matrix.cons_val_zero, LinearMap.add_apply, - LinearMap.smul_apply, innerₛₗ_apply, real_inner_self_eq_norm_sq, smul_eq_mul, - areaForm_apply_self, mul_zero, add_zero, Real.rpow_two, real_inner_comm] + · simp only [Fin.zero_eta, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_zero, LinearMap.add_apply, LinearMap.smul_apply, innerₛₗ_apply, + real_inner_self_eq_norm_sq, smul_eq_mul, areaForm_apply_self, mul_zero, add_zero, + real_inner_comm] ring - · simp only [Fin.mk_one, coe_basisRightAngleRotation, Matrix.cons_val_one, Matrix.head_cons, - LinearMap.add_apply, LinearMap.smul_apply, innerₛₗ_apply, inner_rightAngleRotation_right, - areaForm_apply_self, neg_zero, smul_eq_mul, mul_zero, areaForm_rightAngleRotation_right, - real_inner_self_eq_norm_sq, zero_add, Real.rpow_two, mul_neg] + · simp only [Fin.mk_one, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_one, Matrix.head_cons, LinearMap.add_apply, + LinearMap.smul_apply, innerₛₗ_apply, inner_rightAngleRotation_right, areaForm_apply_self, + neg_zero, smul_eq_mul, mul_zero, areaForm_rightAngleRotation_right, + real_inner_self_eq_norm_sq, zero_add, mul_neg] rw [o.areaForm_swap] ring @@ -381,15 +383,16 @@ theorem inner_mul_areaForm_sub' (a x : E) : ⟪a, x⟫ • ω a - ω a x • inn apply (o.basisRightAngleRotation a ha).ext intro i fin_cases i - · simp only [o.areaForm_swap a x, neg_smul, sub_neg_eq_add, Fin.mk_zero, - coe_basisRightAngleRotation, Matrix.cons_val_zero, LinearMap.add_apply, LinearMap.smul_apply, - areaForm_apply_self, smul_eq_mul, mul_zero, innerₛₗ_apply, real_inner_self_eq_norm_sq, - zero_add, Real.rpow_two] + · simp only [o.areaForm_swap a x, neg_smul, sub_neg_eq_add, Fin.zero_eta, Fin.isValue, id_eq, + coe_basisRightAngleRotation, Nat.succ_eq_add_one, Nat.reduceAdd, Matrix.cons_val_zero, + LinearMap.add_apply, LinearMap.smul_apply, areaForm_apply_self, smul_eq_mul, mul_zero, + innerₛₗ_apply, real_inner_self_eq_norm_sq, zero_add] ring - · simp only [Fin.mk_one, coe_basisRightAngleRotation, Matrix.cons_val_one, Matrix.head_cons, - LinearMap.sub_apply, LinearMap.smul_apply, areaForm_rightAngleRotation_right, - real_inner_self_eq_norm_sq, smul_eq_mul, innerₛₗ_apply, inner_rightAngleRotation_right, - areaForm_apply_self, neg_zero, mul_zero, sub_zero, Real.rpow_two, real_inner_comm] + · simp only [Fin.mk_one, Fin.isValue, id_eq, coe_basisRightAngleRotation, Nat.succ_eq_add_one, + Nat.reduceAdd, Matrix.cons_val_one, Matrix.head_cons, LinearMap.sub_apply, + LinearMap.smul_apply, areaForm_rightAngleRotation_right, real_inner_self_eq_norm_sq, + smul_eq_mul, innerₛₗ_apply, inner_rightAngleRotation_right, areaForm_apply_self, neg_zero, + mul_zero, sub_zero, real_inner_comm] ring /-- For vectors `a x y : E`, the identity `⟪a, x⟫ * ω a y - ω a x * ⟪a, y⟫ = ‖a‖ ^ 2 * ω x y`. -/ diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 03674193faaec..4b86b7218cdfb 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -488,9 +488,17 @@ theorem length_mapAccumr₂ : end MapAccumr +/- #adaptation_note: this attribute should be removed after Mathlib moves to v4.15.0-rc1. -/ +set_option allowUnsafeReducibility true in +attribute [semireducible] Fin.foldr.loop + /-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ -def finRange (n : ℕ) : List (Fin n) := - (range n).pmap Fin.mk fun _ => List.mem_range.1 +-- Note that we use `ofFn (fun x => x)` instead of `ofFn id` to avoid leaving `id` in the terms +-- for e.g. `Fintype (Fin n)`. +def finRange (n : Nat) : List (Fin n) := ofFn (fun x => x) + +-- Verify that `finRange` is semireducible. +example : finRange 3 = [0, 1, 2] := rfl section Deprecated diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index 62d8122999597..bf592cfe78937 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn -import Mathlib.Data.List.Range import Batteries.Data.List.Perm +import Mathlib.Data.List.Nodup /-! # Lists of elements of `Fin n` @@ -21,10 +21,66 @@ namespace List variable {α : Type u} + +theorem finRange_eq_pmap_range (n : ℕ) : finRange n = (range n).pmap Fin.mk (by simp) := by + apply List.ext_getElem <;> simp [finRange] + +@[simp] +theorem finRange_zero : finRange 0 = [] := rfl + +@[simp] +theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := by + rw [finRange_eq_pmap_range] + exact mem_pmap.2 + ⟨a.1, mem_range.2 a.2, by + cases a + rfl⟩ + +theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := by + rw [finRange_eq_pmap_range] + exact (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ + +@[simp] +theorem length_finRange (n : ℕ) : (finRange n).length = n := by + simp [finRange] + +@[simp] +theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by + rw [← length_eq_zero, length_finRange] + +theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := by + rw [finRange_eq_pmap_range] + exact (List.pairwise_lt_range n).pmap (by simp) (by simp) + +theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := by + rw [finRange_eq_pmap_range] + exact (List.pairwise_le_range n).pmap (by simp) (by simp) + +@[simp] +theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : + (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by + simp [finRange, getElem_range, getElem_pmap] + +-- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem +theorem get_finRange {n : ℕ} {i : ℕ} (h) : + (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by + simp + +@[deprecated (since := "2024-08-19")] alias nthLe_finRange := get_finRange + +@[simp] +theorem finRange_map_get (l : List α) : (finRange l.length).map l.get = l := + List.ext_get (by simp) (by simp) + +@[simp] theorem indexOf_finRange {k : ℕ} (i : Fin k) : (finRange k).indexOf i = i := by + have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length.mpr (by simp) + have h₁ : (finRange k).get ⟨(finRange k).indexOf i, this⟩ = i := indexOf_get this + have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ + simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) + @[simp] theorem map_coe_finRange (n : ℕ) : ((finRange n) : List (Fin n)).map (Fin.val) = List.range n := by - simp_rw [finRange, map_pmap, pmap_eq_map] - exact List.map_id _ + apply List.ext_getElem <;> simp theorem finRange_succ_eq_map (n : ℕ) : finRange n.succ = 0 :: (finRange n).map Fin.succ := by apply map_injective_iff.mpr Fin.val_injective @@ -45,7 +101,7 @@ theorem ofFn_eq_pmap {n} {f : Fin n → α} : exact ext_getElem (by simp) fun i hi1 hi2 => by simp [List.getElem_ofFn f i hi1] theorem ofFn_id (n) : ofFn id = finRange n := - ofFn_eq_pmap + rfl theorem ofFn_eq_map {n} {f : Fin n → α} : ofFn f = (finRange n).map f := by rw [← ofFn_id, map_ofFn, Function.comp_id] diff --git a/Mathlib/Data/List/Range.lean b/Mathlib/Data/List/Range.lean index 7cefc522da8e2..e133f0b507d6d 100644 --- a/Mathlib/Data/List/Range.lean +++ b/Mathlib/Data/List/Range.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Kenny Lau, Kim Morrison -/ import Mathlib.Data.List.Chain -import Mathlib.Data.List.Nodup /-! # Ranges of naturals as lists @@ -42,58 +41,10 @@ theorem chain_range_succ (r : ℕ → ℕ → Prop) (n a : ℕ) : rw [range_succ_eq_map, chain_cons, and_congr_right_iff, ← chain'_range_succ, range_succ_eq_map] exact fun _ => Iff.rfl -@[simp] -theorem finRange_zero : finRange 0 = [] := - rfl - -@[simp] -theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := - mem_pmap.2 - ⟨a.1, mem_range.2 a.2, by - cases a - rfl⟩ - -theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := - (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ - -@[simp] -theorem length_finRange (n : ℕ) : (finRange n).length = n := by - rw [finRange, length_pmap, length_range] - -@[simp] -theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by - rw [← length_eq_zero, length_finRange] - -theorem pairwise_lt_finRange (n : ℕ) : Pairwise (· < ·) (finRange n) := - (List.pairwise_lt_range n).pmap (by simp) (by simp) - -theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := - (List.pairwise_le_range n).pmap (by simp) (by simp) - -@[simp] -theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by - simp [finRange, getElem_range, getElem_pmap] - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem -theorem get_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by - simp - @[deprecated (since := "2024-08-19")] alias nthLe_range' := get_range' @[deprecated (since := "2024-08-19")] alias nthLe_range'_1 := getElem_range'_1 @[deprecated (since := "2024-08-19")] alias nthLe_range := get_range -@[deprecated (since := "2024-08-19")] alias nthLe_finRange := get_finRange - -@[simp] -theorem finRange_map_get (l : List α) : (finRange l.length).map l.get = l := - List.ext_get (by simp) (by simp) -@[simp] theorem indexOf_finRange {k : ℕ} (i : Fin k) : (finRange k).indexOf i = i := by - have : (finRange k).indexOf i < (finRange k).length := indexOf_lt_length.mpr (by simp) - have h₁ : (finRange k).get ⟨(finRange k).indexOf i, this⟩ = i := indexOf_get this - have h₂ : (finRange k).get ⟨i, by simp⟩ = i := get_finRange _ - simpa using (Nodup.get_inj_iff (nodup_finRange k)).mp (Eq.trans h₁ h₂.symm) section Ranges diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index cc89adca8c9e8..9a2c8391406a8 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Choose.Basic -import Mathlib.Data.List.Range +import Mathlib.Data.List.FinRange import Mathlib.Data.List.Perm.Basic +import Mathlib.Data.List.Lex /-! # sublists diff --git a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean index 9476da73472ed..308268392107c 100644 --- a/Mathlib/LinearAlgebra/AffineSpace/Independent.lean +++ b/Mathlib/LinearAlgebra/AffineSpace/Independent.lean @@ -618,7 +618,7 @@ theorem affineIndependent_of_ne {p₁ p₂ : P} (h : p₁ ≠ p₂) : AffineInde ext fin_cases i · simp at hi - · simp only [Fin.val_one] + · simp haveI : Unique { x // x ≠ (0 : Fin 2) } := ⟨⟨i₁⟩, he'⟩ apply linearIndependent_unique rw [he' default] diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 7eb5dbb5ffec4..59b083317b35b 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -207,17 +207,19 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : ext ⟨g, rfl⟩ i j : 3 fin_cases i <;> [fin_cases j; skip] -- the following are proved by `simp`, but it is replaced by `simp only` to avoid timeouts. - · simp only [mB, mulVec, dotProduct, Fin.sum_univ_two, coe_matrix_coe, - Int.coe_castRingHom, lcRow0_apply, Function.comp_apply, cons_val_zero, lcRow0Extend_apply, + · simp only [Fin.isValue, Int.cast_one, map_apply_coe, RingHom.mapMatrix_apply, + Int.coe_castRingHom, lcRow0_apply, map_apply, Fin.zero_eta, id_eq, Function.comp_apply, + of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, lcRow0Extend_apply, LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, - val_planeConformalMatrix, neg_neg, mulVecLin_apply, cons_val_one, head_cons, of_apply, - Fin.mk_zero, Fin.mk_one] + val_planeConformalMatrix, neg_neg, mulVecLin_apply, mulVec, dotProduct, Fin.sum_univ_two, + cons_val_one, head_cons, mB, f₁] · convert congr_arg (fun n : ℤ => (-n : ℝ)) g.det_coe.symm using 1 - simp only [f₁, mulVec, dotProduct, Fin.sum_univ_two, Matrix.det_fin_two, Function.comp_apply, - Subtype.coe_mk, lcRow0Extend_apply, cons_val_zero, + simp only [Fin.zero_eta, id_eq, Function.comp_apply, lcRow0Extend_apply, cons_val_zero, LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, - val_planeConformalMatrix, mulVecLin_apply, cons_val_one, head_cons, map_apply, neg_mul, - Int.cast_sub, Int.cast_mul, neg_sub, of_apply, Fin.mk_zero, Fin.mk_one] + mulVecLin_apply, mulVec, dotProduct, det_fin_two, f₁] + simp only [Fin.isValue, Fin.mk_one, val_planeConformalMatrix, neg_neg, of_apply, cons_val', + empty_val', cons_val_fin_one, cons_val_one, head_fin_const, map_apply, Fin.sum_univ_two, + cons_val_zero, neg_mul, head_cons, Int.cast_sub, Int.cast_mul, neg_sub] ring · rfl diff --git a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean index 84f4e2aa14e91..3b10a1368dea7 100644 --- a/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean +++ b/Mathlib/NumberTheory/NumberField/CanonicalEmbedding/Basic.lean @@ -585,17 +585,17 @@ theorem stdBasis_repr_eq_matrixToStdBasis_mul (x : (K →+* ℂ) → ℂ) | inr c => rcases c with ⟨w, j⟩ fin_cases j - · simp_rw [Fin.mk_zero, stdBasis_apply_ofIsComplex_fst, fromBlocks_apply₂₁, - fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, - blockDiagonal_apply, Prod.swap_prod_mk, ite_mul, zero_mul, sum_const_zero, zero_add, - sum_add_distrib, sum_ite_eq, mem_univ, ite_true, of_apply, cons_val', cons_val_zero, - cons_val_one, head_cons, ← hx (embedding w), re_eq_add_conj] + · simp only [Fin.zero_eta, Fin.isValue, id_eq, stdBasis_apply_ofIsComplex_fst, re_eq_add_conj, + mul_neg, fromBlocks_apply₂₁, zero_apply, zero_mul, sum_const_zero, fromBlocks_apply₂₂, + submatrix_apply, Prod.swap_prod_mk, blockDiagonal_apply, of_apply, cons_val', cons_val_zero, + empty_val', cons_val_fin_one, ite_mul, cons_val_one, head_cons, sum_add_distrib, sum_ite_eq, + mem_univ, ↓reduceIte, ← hx (embedding w), zero_add] field_simp - · simp_rw [Fin.mk_one, stdBasis_apply_ofIsComplex_snd, fromBlocks_apply₂₁, - fromBlocks_apply₂₂, Matrix.zero_apply, submatrix_apply, blockDiagonal_apply, - Prod.swap_prod_mk, ite_mul, zero_mul, sum_const_zero, zero_add, sum_add_distrib, sum_ite_eq, - mem_univ, ite_true, of_apply, cons_val', cons_val_zero, cons_val_one, head_cons, - ← hx (embedding w), im_eq_sub_conj] + · simp only [Fin.mk_one, Fin.isValue, id_eq, stdBasis_apply_ofIsComplex_snd, im_eq_sub_conj, + mul_neg, fromBlocks_apply₂₁, zero_apply, zero_mul, sum_const_zero, fromBlocks_apply₂₂, + submatrix_apply, Prod.swap_prod_mk, blockDiagonal_apply, of_apply, cons_val', cons_val_zero, + empty_val', cons_val_fin_one, cons_val_one, head_fin_const, ite_mul, neg_mul, head_cons, + sum_add_distrib, sum_ite_eq, mem_univ, ↓reduceIte, ← hx (embedding w), zero_add] ring_nf; field_simp end stdBasis diff --git a/Mathlib/Order/RelSeries.lean b/Mathlib/Order/RelSeries.lean index b2c41395b294b..e385b759958c2 100644 --- a/Mathlib/Order/RelSeries.lean +++ b/Mathlib/Order/RelSeries.lean @@ -112,9 +112,9 @@ corresponds to each other. -/ protected def Equiv : RelSeries r ≃ {x : List α | x ≠ [] ∧ x.Chain' r} where toFun x := ⟨_, x.toList_ne_nil, x.toList_chain'⟩ invFun x := fromListChain' _ x.2.1 x.2.2 - left_inv x := ext (by simp [toList]) <| by ext; apply List.get_ofFn + left_inv x := ext (by simp [toList]) <| by ext; dsimp; apply List.get_ofFn right_inv x := by - refine Subtype.ext (List.ext_get ?_ fun n hn1 _ => List.get_ofFn _ _) + refine Subtype.ext (List.ext_get ?_ fun n hn1 _ => by dsimp; apply List.get_ofFn) have := Nat.succ_pred_eq_of_pos <| List.length_pos.mpr x.2.1 simp_all [toList] diff --git a/Mathlib/RingTheory/Complex.lean b/Mathlib/RingTheory/Complex.lean index 68be8e6b337a3..1f3b98e14ae26 100644 --- a/Mathlib/RingTheory/Complex.lean +++ b/Mathlib/RingTheory/Complex.lean @@ -18,11 +18,11 @@ theorem Algebra.leftMulMatrix_complex (z : ℂ) : rw [Algebra.leftMulMatrix_eq_repr_mul, Complex.coe_basisOneI_repr, Complex.coe_basisOneI, mul_re, mul_im, Matrix.of_apply] fin_cases j - · simp only [Fin.mk_zero, Matrix.cons_val_zero, one_re, mul_one, one_im, mul_zero, sub_zero, - zero_add] + · simp only [Fin.zero_eta, id_eq, Matrix.cons_val_zero, one_re, mul_one, one_im, mul_zero, + sub_zero, zero_add, Matrix.cons_val_fin_one] fin_cases i <;> rfl - · simp only [Fin.mk_one, Matrix.cons_val_one, Matrix.head_cons, I_re, mul_zero, I_im, mul_one, - zero_sub, add_zero] + · simp only [Fin.mk_one, id_eq, Matrix.cons_val_one, Matrix.head_cons, I_re, mul_zero, I_im, + mul_one, zero_sub, add_zero, Matrix.cons_val_fin_one] fin_cases i <;> rfl theorem Algebra.trace_complex_apply (z : ℂ) : Algebra.trace ℝ ℂ z = 2 * z.re := by diff --git a/Mathlib/SetTheory/Cardinal/Subfield.lean b/Mathlib/SetTheory/Cardinal/Subfield.lean index 604447735f1cc..37f45168df1dd 100644 --- a/Mathlib/SetTheory/Cardinal/Subfield.lean +++ b/Mathlib/SetTheory/Cardinal/Subfield.lean @@ -76,7 +76,7 @@ lemma cardinalMk_closure_le_max : #(closure s) ≤ max #s ℵ₀ := exact (add_lt_aleph0 this h).le · rw [max_eq_left h, add_eq_right h (this.le.trans h), max_eq_left h] rintro (n|_) - · fin_cases n <;> infer_instance + · fin_cases n <;> (dsimp only [id_eq]; infer_instance) infer_instance @[deprecated (since := "2024-11-10")] alias cardinal_mk_closure_le_max := cardinalMk_closure_le_max From 3e56ff8fc708e700c810326438a65f57315ec0e9 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Mon, 25 Nov 2024 15:15:47 +0000 Subject: [PATCH 146/250] feat(CategoryTheory): limits of eventually constant functors from cofiltered categories (#19021) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit An eventually constant functor from a cofiltered category admits a limit. Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Filtered/Basic.lean | 20 ++ .../Constructions/EventuallyConstant.lean | 260 ++++++++++++++++++ 3 files changed, 281 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean diff --git a/Mathlib.lean b/Mathlib.lean index 39dccdce392ff..02bb3ae68c288 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1770,6 +1770,7 @@ import Mathlib.CategoryTheory.Limits.Connected import Mathlib.CategoryTheory.Limits.Constructions.BinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.EpiMono import Mathlib.CategoryTheory.Limits.Constructions.Equalizers +import Mathlib.CategoryTheory.Limits.Constructions.EventuallyConstant import Mathlib.CategoryTheory.Limits.Constructions.Filtered import Mathlib.CategoryTheory.Limits.Constructions.FiniteProductsOfBinaryProducts import Mathlib.CategoryTheory.Limits.Constructions.LimitsOfProductsAndEqualizers diff --git a/Mathlib/CategoryTheory/Filtered/Basic.lean b/Mathlib/CategoryTheory/Filtered/Basic.lean index 28de77ac9e3bb..584411d3326a4 100644 --- a/Mathlib/CategoryTheory/Filtered/Basic.lean +++ b/Mathlib/CategoryTheory/Filtered/Basic.lean @@ -616,6 +616,26 @@ theorem _root_.CategoryTheory.Functor.ranges_directed (F : C ⥤ Type*) (j : C) let ⟨l, li, lk, e⟩ := cospan ij kj refine ⟨⟨l, lk ≫ kj⟩, e ▸ ?_, ?_⟩ <;> simp_rw [F.map_comp] <;> apply Set.range_comp_subset_range +/-- Given a "bowtie" of morphisms +``` + k₁ k₂ + |\ /| + | \/ | + | /\ | + |/ \∣ + vv vv + j₁ j₂ +``` +in a cofiltered category, we can construct an object `s` and two morphisms +from `s` to `k₁` and `k₂`, making the resulting squares commute. +-/ +theorem bowtie {j₁ j₂ k₁ k₂ : C} (f₁ : k₁ ⟶ j₁) (g₁ : k₂ ⟶ j₁) (f₂ : k₁ ⟶ j₂) (g₂ : k₂ ⟶ j₂) : + ∃ (s : C) (α : s ⟶ k₁) (β : s ⟶ k₂), α ≫ f₁ = β ≫ g₁ ∧ α ≫ f₂ = β ≫ g₂ := by + obtain ⟨t, k₁t, k₂t, ht⟩ := cospan f₁ g₁ + obtain ⟨s, ts, hs⟩ := IsCofilteredOrEmpty.cone_maps (k₁t ≫ f₂) (k₂t ≫ g₂) + exact ⟨s, ts ≫ k₁t, ts ≫ k₂t, by simp only [Category.assoc, ht], + by simp only [Category.assoc, hs]⟩ + end AllowEmpty end IsCofiltered diff --git a/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean b/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean new file mode 100644 index 0000000000000..9ae75d85d62ae --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Constructions/EventuallyConstant.lean @@ -0,0 +1,260 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ +import Mathlib.CategoryTheory.Filtered.Basic +import Mathlib.CategoryTheory.Limits.HasLimits + +/-! +# Limits of eventually constant functors + +If `F : J ⥤ C` is a functor from a cofiltered category, and `j : J`, +we introduce a property `F.IsEventuallyConstantTo j` which says +that for any `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +Under this assumption, it is shown that `F` admits `F.obj j` as a limit +(`Functor.IsEventuallyConstantTo.isLimitCone`). + +A typeclass `Cofiltered.IsEventuallyConstant` is also introduced, and +the dual results for filtered categories and colimits are also obtained. + +-/ + +namespace CategoryTheory + +open Category Limits + +variable {J C : Type*} [Category J] [Category C] (F : J ⥤ C) + +namespace Functor + +/-- A functor `F : J ⥤ C` is eventually constant to `j : J` if +for any map `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +If `J` is cofiltered, this implies `F` has a limit. -/ +def IsEventuallyConstantTo (j : J) : Prop := + ∀ ⦃i : J⦄ (f : i ⟶ j), IsIso (F.map f) + +/-- A functor `F : J ⥤ C` is eventually constant from `i : J` if +for any map `f : i ⟶ j`, the induced morphism `F.map f` is an isomorphism. +If `J` is filtered, this implies `F` has a colimit. -/ +def IsEventuallyConstantFrom (i : J) : Prop := + ∀ ⦃j : J⦄ (f : i ⟶ j), IsIso (F.map f) + +namespace IsEventuallyConstantTo + +variable {F} {i₀ : J} (h : F.IsEventuallyConstantTo i₀) + +include h + +lemma isIso_map {i j : J} (φ : i ⟶ j) (π : j ⟶ i₀) : IsIso (F.map φ) := by + have := h π + have := h (φ ≫ π) + exact IsIso.of_isIso_fac_right (F.map_comp φ π).symm + +lemma precomp {j : J} (f : j ⟶ i₀) : F.IsEventuallyConstantTo j := + fun _ φ ↦ h.isIso_map φ f + +section + +variable {i j : J} (φ : i ⟶ j) (hφ : Nonempty (j ⟶ i₀)) + +/-- The isomorphism `F.obj i ≅ F.obj j` induced by `φ : i ⟶ j`, +when `h : F.IsEventuallyConstantTo i₀` and there exists a map `j ⟶ i₀`. -/ +@[simps! hom] +noncomputable def isoMap : F.obj i ≅ F.obj j := + have := h.isIso_map φ hφ.some + asIso (F.map φ) + +@[reassoc (attr := simp)] +lemma isoMap_hom_inv_id : F.map φ ≫ (h.isoMap φ hφ).inv = 𝟙 _ := + (h.isoMap φ hφ).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoMap_inv_hom_id : (h.isoMap φ hφ).inv ≫ F.map φ = 𝟙 _ := + (h.isoMap φ hφ).inv_hom_id + +end + +variable [IsCofiltered J] +open IsCofiltered + +/-- Auxiliary definition for `IsEventuallyConstantTo.cone`. -/ +noncomputable def coneπApp (j : J) : F.obj i₀ ⟶ F.obj j := + (h.isoMap (minToLeft i₀ j) ⟨𝟙 _⟩).inv ≫ F.map (minToRight i₀ j) + +lemma coneπApp_eq (j j' : J) (α : j' ⟶ i₀) (β : j' ⟶ j) : + h.coneπApp j = (h.isoMap α ⟨𝟙 _⟩).inv ≫ F.map β := by + obtain ⟨s, γ, δ, h₁, h₂⟩ := IsCofiltered.bowtie + (IsCofiltered.minToRight i₀ j) β (IsCofiltered.minToLeft i₀ j) α + dsimp [coneπApp] + rw [← cancel_epi ((h.isoMap α ⟨𝟙 _⟩).hom), isoMap_hom, isoMap_hom_inv_id_assoc, + ← cancel_epi (h.isoMap δ ⟨α⟩).hom, isoMap_hom, + ← F.map_comp δ β, ← h₁, F.map_comp, ← F.map_comp_assoc, ← h₂, F.map_comp_assoc, + isoMap_hom_inv_id_assoc] + +@[simp] +lemma coneπApp_eq_id : h.coneπApp i₀ = 𝟙 _ := by + rw [h.coneπApp_eq i₀ i₀ (𝟙 _) (𝟙 _), h.isoMap_inv_hom_id] + +/-- Given `h : F.IsEventuallyConstantTo i₀`, this is the (limit) cone for `F` whose +point is `F.obj i₀`. -/ +@[simps] +noncomputable def cone : Cone F where + pt := F.obj i₀ + π := + { app := h.coneπApp + naturality := fun j j' φ ↦ by + dsimp + rw [id_comp] + let i := IsCofiltered.min i₀ j + let α : i ⟶ i₀ := IsCofiltered.minToLeft _ _ + let β : i ⟶ j := IsCofiltered.minToRight _ _ + rw [h.coneπApp_eq j _ α β, assoc, h.coneπApp_eq j' _ α (β ≫ φ), map_comp] } + +/-- When `h : F.IsEventuallyConstantTo i₀`, the limit of `F` exists and is `F.obj i₀`. -/ +def isLimitCone : IsLimit h.cone where + lift s := s.π.app i₀ + fac s j := by + dsimp [coneπApp] + rw [← s.w (IsCofiltered.minToLeft i₀ j), ← s.w (IsCofiltered.minToRight i₀ j), assoc, + isoMap_hom_inv_id_assoc] + uniq s m hm := by simp only [← hm i₀, cone_π_app, coneπApp_eq_id, cone_pt, comp_id] + +lemma hasLimit : HasLimit F := ⟨_, h.isLimitCone⟩ + +lemma isIso_π_of_isLimit {c : Cone F} (hc : IsLimit c) : + IsIso (c.π.app i₀) := by + simp only [← IsLimit.conePointUniqueUpToIso_hom_comp hc h.isLimitCone i₀, + cone_π_app, coneπApp_eq_id, cone_pt, comp_id] + infer_instance + +/-- More general version of `isIso_π_of_isLimit`. -/ +lemma isIso_π_of_isLimit' {c : Cone F} (hc : IsLimit c) (j : J) (π : j ⟶ i₀) : + IsIso (c.π.app j) := + (h.precomp π).isIso_π_of_isLimit hc + +end IsEventuallyConstantTo + +namespace IsEventuallyConstantFrom + +variable {F} {i₀ : J} (h : F.IsEventuallyConstantFrom i₀) + +include h + +lemma isIso_map {i j : J} (φ : i ⟶ j) (ι : i₀ ⟶ i) : IsIso (F.map φ) := by + have := h ι + have := h (ι ≫ φ) + exact IsIso.of_isIso_fac_left (F.map_comp ι φ).symm + +lemma postcomp {j : J} (f : i₀ ⟶ j) : F.IsEventuallyConstantFrom j := + fun _ φ ↦ h.isIso_map φ f + +section + +variable {i j : J} (φ : i ⟶ j) (hφ : Nonempty (i₀ ⟶ i)) + +/-- The isomorphism `F.obj i ≅ F.obj j` induced by `φ : i ⟶ j`, +when `h : F.IsEventuallyConstantFrom i₀` and there exists a map `i₀ ⟶ i`. -/ +@[simps! hom] +noncomputable def isoMap : F.obj i ≅ F.obj j := + have := h.isIso_map φ hφ.some + asIso (F.map φ) + +@[reassoc (attr := simp)] +lemma isoMap_hom_inv_id : F.map φ ≫ (h.isoMap φ hφ).inv = 𝟙 _ := + (h.isoMap φ hφ).hom_inv_id + +@[reassoc (attr := simp)] +lemma isoMap_inv_hom_id : (h.isoMap φ hφ).inv ≫ F.map φ = 𝟙 _ := + (h.isoMap φ hφ).inv_hom_id + +end + +variable [IsFiltered J] +open IsFiltered + +/-- Auxiliary definition for `IsEventuallyConstantFrom.cocone`. -/ +noncomputable def coconeιApp (j : J) : F.obj j ⟶ F.obj i₀ := + F.map (rightToMax i₀ j) ≫ (h.isoMap (leftToMax i₀ j) ⟨𝟙 _⟩).inv + +lemma coconeιApp_eq (j j' : J) (α : j ⟶ j') (β : i₀ ⟶ j') : + h.coconeιApp j = F.map α ≫ (h.isoMap β ⟨𝟙 _⟩).inv := by + obtain ⟨s, γ, δ, h₁, h₂⟩ := IsFiltered.bowtie + (IsFiltered.leftToMax i₀ j) β (IsFiltered.rightToMax i₀ j) α + dsimp [coconeιApp] + rw [← cancel_mono ((h.isoMap β ⟨𝟙 _⟩).hom), assoc, assoc, isoMap_hom, isoMap_inv_hom_id, + comp_id, ← cancel_mono (h.isoMap δ ⟨β⟩).hom, isoMap_hom, assoc, assoc, ← F.map_comp α δ, + ← h₂, F.map_comp, ← F.map_comp β δ, ← h₁, F.map_comp, isoMap_inv_hom_id_assoc] + +@[simp] +lemma coconeιApp_eq_id : h.coconeιApp i₀ = 𝟙 _ := by + rw [h.coconeιApp_eq i₀ i₀ (𝟙 _) (𝟙 _), h.isoMap_hom_inv_id] + +/-- Given `h : F.IsEventuallyConstantFrom i₀`, this is the (limit) cocone for `F` whose +point is `F.obj i₀`. -/ +@[simps] +noncomputable def cocone : Cocone F where + pt := F.obj i₀ + ι := + { app := h.coconeιApp + naturality := fun j j' φ ↦ by + dsimp + rw [comp_id] + let i := IsFiltered.max i₀ j' + let α : i₀ ⟶ i := IsFiltered.leftToMax _ _ + let β : j' ⟶ i := IsFiltered.rightToMax _ _ + rw [h.coconeιApp_eq j' _ β α, h.coconeιApp_eq j _ (φ ≫ β) α, map_comp, assoc] } + +/-- When `h : F.IsEventuallyConstantFrom i₀`, the colimit of `F` exists and is `F.obj i₀`. -/ +def isColimitCocone : IsColimit h.cocone where + desc s := s.ι.app i₀ + fac s j := by + dsimp [coconeιApp] + rw [← s.w (IsFiltered.rightToMax i₀ j), ← s.w (IsFiltered.leftToMax i₀ j), assoc, + isoMap_inv_hom_id_assoc] + uniq s m hm := by simp only [← hm i₀, cocone_ι_app, coconeιApp_eq_id, id_comp] + +lemma hasColimit : HasColimit F := ⟨_, h.isColimitCocone⟩ + +lemma isIso_ι_of_isColimit {c : Cocone F} (hc : IsColimit c) : + IsIso (c.ι.app i₀) := by + simp only [← IsColimit.comp_coconePointUniqueUpToIso_inv hc h.isColimitCocone i₀, + cocone_ι_app, coconeιApp_eq_id, id_comp] + infer_instance + +/-- More general version of `isIso_ι_of_isColimit`. -/ +lemma isIso_ι_of_isColimit' {c : Cocone F} (hc : IsColimit c) (j : J) (ι : i₀ ⟶ j) : + IsIso (c.ι.app j) := + (h.postcomp ι).isIso_ι_of_isColimit hc + +end IsEventuallyConstantFrom + +end Functor + +namespace IsCofiltered + +/-- A functor `F : J ⥤ C` from a cofiltered category is eventually constant if there +exists `j : J`, such that for any `f : i ⟶ j`, the induced map `F.map f` is an isomorphism. -/ +class IsEventuallyConstant : Prop where + exists_isEventuallyConstantTo : ∃ (j : J), F.IsEventuallyConstantTo j + +instance [hF : IsEventuallyConstant F] [IsCofiltered J] : HasLimit F := by + obtain ⟨j, h⟩ := hF.exists_isEventuallyConstantTo + exact h.hasLimit + +end IsCofiltered + +namespace IsFiltered + +/-- A functor `F : J ⥤ C` from a filtered category is eventually constant if there +exists `i : J`, such that for any `f : i ⟶ j`, the induced map `F.map f` is an isomorphism. -/ +class IsEventuallyConstant : Prop where + exists_isEventuallyConstantFrom : ∃ (i : J), F.IsEventuallyConstantFrom i + +instance [hF : IsEventuallyConstant F] [IsFiltered J] : HasColimit F := by + obtain ⟨j, h⟩ := hF.exists_isEventuallyConstantFrom + exact h.hasColimit + +end IsFiltered + +end CategoryTheory From b9e31e9b11f73ac54f481c8fe732f94d825895a4 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Mon, 25 Nov 2024 15:15:49 +0000 Subject: [PATCH 147/250] feat(CategoryTheory): final functors and preservation/reflection/creation of colimits (#19452) Co-authored-by: Markus Himmel --- Mathlib/CategoryTheory/Limits/Final.lean | 126 ++++++++++++++++++++++- 1 file changed, 125 insertions(+), 1 deletion(-) diff --git a/Mathlib/CategoryTheory/Limits/Final.lean b/Mathlib/CategoryTheory/Limits/Final.lean index 69444f0842058..6719a13bc510a 100644 --- a/Mathlib/CategoryTheory/Limits/Final.lean +++ b/Mathlib/CategoryTheory/Limits/Final.lean @@ -65,7 +65,7 @@ Dualise condition 3 above and the implications 2 ⇒ 3 and 3 ⇒ 1 to initial fu noncomputable section -universe v v₁ v₂ v₃ u₁ u₂ u₃ +universe v v₁ v₂ v₃ v₄ u₁ u₂ u₃ u₄ namespace CategoryTheory @@ -305,6 +305,27 @@ def colimitCoconeComp (t : ColimitCocone G) : ColimitCocone (F ⋙ G) where instance (priority := 100) comp_hasColimit [HasColimit G] : HasColimit (F ⋙ G) := HasColimit.mk (colimitCoconeComp F (getColimitCocone G)) +instance (priority := 100) comp_preservesColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesColimit G H] : PreservesColimit (F ⋙ G) H where + preserves {c} hc := by + refine ⟨isColimitExtendCoconeEquiv (G := G ⋙ H) F (H.mapCocone c) ?_⟩ + let hc' := isColimitOfPreserves H ((isColimitExtendCoconeEquiv F c).symm hc) + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) comp_reflectsColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsColimit G H] : ReflectsColimit (F ⋙ G) H where + reflects {c} hc := by + refine ⟨isColimitExtendCoconeEquiv F _ (isColimitOfReflects H ?_)⟩ + let hc' := (isColimitExtendCoconeEquiv (G := G ⋙ H) F _).symm hc + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) compCreatesColimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesColimit G H] : CreatesColimit (F ⋙ G) H where + lifts {c} hc := by + refine ⟨(liftColimit ((isColimitExtendCoconeEquiv F (G := G ⋙ H) _).symm hc)).whisker F, ?_⟩ + let i := liftedColimitMapsToOriginal ((isColimitExtendCoconeEquiv F (G := G ⋙ H) _).symm hc) + exact (Cocones.whiskering F).mapIso i ≪≫ ((coconesEquiv F (G ⋙ H)).unitIso.app _).symm + instance colimit_pre_isIso [HasColimit G] : IsIso (colimit.pre G F) := by rw [colimit.pre_eq (colimitCoconeComp F (getColimitCocone G)) (getColimitCocone G)] erw [IsColimit.desc_self] @@ -361,10 +382,51 @@ We can't make this an instance, because `F` is not determined by the goal. theorem hasColimit_of_comp [HasColimit (F ⋙ G)] : HasColimit G := HasColimit.mk (colimitCoconeOfComp F (getColimitCocone (F ⋙ G))) +theorem preservesColimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesColimit (F ⋙ G) H] : PreservesColimit G H where + preserves {c} hc := by + refine ⟨isColimitWhiskerEquiv F _ ?_⟩ + let hc' := isColimitOfPreserves H ((isColimitWhiskerEquiv F _).symm hc) + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +theorem reflectsColimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsColimit (F ⋙ G) H] : ReflectsColimit G H where + reflects {c} hc := by + refine ⟨isColimitWhiskerEquiv F _ (isColimitOfReflects H ?_)⟩ + let hc' := (isColimitWhiskerEquiv F _).symm hc + exact IsColimit.ofIsoColimit hc' (Cocones.ext (Iso.refl _) (by simp)) + +/-- If `F` is final and `F ⋙ G` creates colimits of `H`, then so does `G`. -/ +def createsColimitOfComp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesColimit (F ⋙ G) H] : CreatesColimit G H where + reflects := (reflectsColimit_of_comp F).reflects + lifts {c} hc := by + refine ⟨(extendCocone (F := F)).obj (liftColimit ((isColimitWhiskerEquiv F _).symm hc)), ?_⟩ + let i := liftedColimitMapsToOriginal (K := (F ⋙ G)) ((isColimitWhiskerEquiv F _).symm hc) + refine ?_ ≪≫ ((extendCocone (F := F)).mapIso i) ≪≫ ((coconesEquiv F (G ⋙ H)).counitIso.app _) + exact Cocones.ext (Iso.refl _) + include F in theorem hasColimitsOfShape_of_final [HasColimitsOfShape C E] : HasColimitsOfShape D E where has_colimit := fun _ => hasColimit_of_comp F +include F in +theorem preservesColimitsOfShape_of_final {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [PreservesColimitsOfShape C H] : PreservesColimitsOfShape D H where + preservesColimit := preservesColimit_of_comp F + +include F in +theorem reflectsColimitsOfShape_of_final {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [ReflectsColimitsOfShape C H] : ReflectsColimitsOfShape D H where + reflectsColimit := reflectsColimit_of_comp F + +include F in +/-- If `H` creates colimits of shape `C` and `F : C ⥤ D` is final, then `H` creates colimits of +shape `D`. -/ +def createsColimitsOfShapeOfFinal {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [CreatesColimitsOfShape C H] : CreatesColimitsOfShape D H where + CreatesColimit := createsColimitOfComp F + end Final end ArbitraryUniverse @@ -592,6 +654,27 @@ def limitConeComp (t : LimitCone G) : LimitCone (F ⋙ G) where instance (priority := 100) comp_hasLimit [HasLimit G] : HasLimit (F ⋙ G) := HasLimit.mk (limitConeComp F (getLimitCone G)) +instance (priority := 100) comp_preservesLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesLimit G H] : PreservesLimit (F ⋙ G) H where + preserves {c} hc := by + refine ⟨isLimitExtendConeEquiv (G := G ⋙ H) F (H.mapCone c) ?_⟩ + let hc' := isLimitOfPreserves H ((isLimitExtendConeEquiv F c).symm hc) + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) comp_reflectsLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsLimit G H] : ReflectsLimit (F ⋙ G) H where + reflects {c} hc := by + refine ⟨isLimitExtendConeEquiv F _ (isLimitOfReflects H ?_)⟩ + let hc' := (isLimitExtendConeEquiv (G := G ⋙ H) F _).symm hc + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +instance (priority := 100) compCreatesLimit {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesLimit G H] : CreatesLimit (F ⋙ G) H where + lifts {c} hc := by + refine ⟨(liftLimit ((isLimitExtendConeEquiv F (G := G ⋙ H) _).symm hc)).whisker F, ?_⟩ + let i := liftedLimitMapsToOriginal ((isLimitExtendConeEquiv F (G := G ⋙ H) _).symm hc) + exact (Cones.whiskering F).mapIso i ≪≫ ((conesEquiv F (G ⋙ H)).unitIso.app _).symm + instance limit_pre_isIso [HasLimit G] : IsIso (limit.pre G F) := by rw [limit.pre_eq (limitConeComp F (getLimitCone G)) (getLimitCone G)] erw [IsLimit.lift_self] @@ -637,10 +720,51 @@ We can't make this an instance, because `F` is not determined by the goal. theorem hasLimit_of_comp [HasLimit (F ⋙ G)] : HasLimit G := HasLimit.mk (limitConeOfComp F (getLimitCone (F ⋙ G))) +theorem preservesLimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [PreservesLimit (F ⋙ G) H] : PreservesLimit G H where + preserves {c} hc := by + refine ⟨isLimitWhiskerEquiv F _ ?_⟩ + let hc' := isLimitOfPreserves H ((isLimitWhiskerEquiv F _).symm hc) + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +theorem reflectsLimit_of_comp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [ReflectsLimit (F ⋙ G) H] : ReflectsLimit G H where + reflects {c} hc := by + refine ⟨isLimitWhiskerEquiv F _ (isLimitOfReflects H ?_)⟩ + let hc' := (isLimitWhiskerEquiv F _).symm hc + exact IsLimit.ofIsoLimit hc' (Cones.ext (Iso.refl _) (by simp)) + +/-- If `F` is initial and `F ⋙ G` creates limits of `H`, then so does `G`. -/ +def createsLimitOfComp {B : Type u₄} [Category.{v₄} B] {H : E ⥤ B} + [CreatesLimit (F ⋙ G) H] : CreatesLimit G H where + reflects := (reflectsLimit_of_comp F).reflects + lifts {c} hc := by + refine ⟨(extendCone (F := F)).obj (liftLimit ((isLimitWhiskerEquiv F _).symm hc)), ?_⟩ + let i := liftedLimitMapsToOriginal (K := (F ⋙ G)) ((isLimitWhiskerEquiv F _).symm hc) + refine ?_ ≪≫ ((extendCone (F := F)).mapIso i) ≪≫ ((conesEquiv F (G ⋙ H)).counitIso.app _) + exact Cones.ext (Iso.refl _) + include F in theorem hasLimitsOfShape_of_initial [HasLimitsOfShape C E] : HasLimitsOfShape D E where has_limit := fun _ => hasLimit_of_comp F +include F in +theorem preservesLimitsOfShape_of_initial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [PreservesLimitsOfShape C H] : PreservesLimitsOfShape D H where + preservesLimit := preservesLimit_of_comp F + +include F in +theorem reflectsLimitsOfShape_of_initial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [ReflectsLimitsOfShape C H] : ReflectsLimitsOfShape D H where + reflectsLimit := reflectsLimit_of_comp F + +include F in +/-- If `H` creates limits of shape `C` and `F : C ⥤ D` is initial, then `H` creates limits of shape +`D`. -/ +def createsLimitsOfShapeOfInitial {B : Type u₄} [Category.{v₄} B] (H : E ⥤ B) + [CreatesLimitsOfShape C H] : CreatesLimitsOfShape D H where + CreatesLimit := createsLimitOfComp F + end Initial section From dbcafd024242e06f4dbf2169991776331edca4a6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 15:32:00 +0000 Subject: [PATCH 148/250] feat: vertical line test for functions, group homs, linear maps (#18822) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Prove the vertical line test for monoid homomorphisms/isomorphisms. Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. Furthermore, if `f` is also surjective on the second factor and its image intersects every "horizontal line" `{(h, i) | h : H}` at most once, then `f'` is actually an isomorphism `f' : H ≃ I`. Co-authored-by: David Loeffler --- Mathlib.lean | 1 + Mathlib/Algebra/Group/Graph.lean | 231 +++++++++++++++++++++++++++++++ Mathlib/Algebra/Group/Prod.lean | 2 + Mathlib/Data/Prod/Basic.lean | 5 +- Mathlib/Data/Set/Function.lean | 51 +++++++ Mathlib/LinearAlgebra/Prod.lean | 82 +++++++++++ 6 files changed, 371 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Algebra/Group/Graph.lean diff --git a/Mathlib.lean b/Mathlib.lean index 02bb3ae68c288..979ce54ccf805 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -271,6 +271,7 @@ import Mathlib.Algebra.Group.Fin.Basic import Mathlib.Algebra.Group.Fin.Tuple import Mathlib.Algebra.Group.FiniteSupport import Mathlib.Algebra.Group.ForwardDiff +import Mathlib.Algebra.Group.Graph import Mathlib.Algebra.Group.Hom.Basic import Mathlib.Algebra.Group.Hom.CompTypeclasses import Mathlib.Algebra.Group.Hom.Defs diff --git a/Mathlib/Algebra/Group/Graph.lean b/Mathlib/Algebra/Group/Graph.lean new file mode 100644 index 0000000000000..a3cf4bee29867 --- /dev/null +++ b/Mathlib/Algebra/Group/Graph.lean @@ -0,0 +1,231 @@ +/- +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, David Loeffler +-/ +import Mathlib.Algebra.Group.Subgroup.Ker + +/-! +# Vertical line test for group homs + +This file proves the vertical line test for monoid homomorphisms/isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. + +Furthermore, if `f` is also surjective on the second factor and its image intersects every +"horizontal line" `{(h, i) | h : H}` at most once, then `f'` is actually an isomorphism +`f' : H ≃ I`. + +We also prove specialised versions when `f` is the inclusion of a subgroup of the direct product. +(The version for general homomorphisms can easily be reduced to this special case, but the +homomorphism version is more flexible in applications.) +-/ + +open Function Set + +variable {G H I : Type*} + +section Monoid +variable [Monoid G] [Monoid H] [Monoid I] + +namespace MonoidHom + +/-- The graph of a monoid homomorphism as a submonoid. + +See also `MonoidHom.graph` for the graph as a subgroup. -/ +@[to_additive +"The graph of a monoid homomorphism as a submonoid. + +See also `AddMonoidHom.graph` for the graph as a subgroup."] +def mgraph (f : G →* H) : Submonoid (G × H) where + carrier := {x | f x.1 = x.2} + one_mem' := map_one f + mul_mem' {x y} := by simp +contextual + +-- TODO: Can `to_additive` be smarter about `simps`? +attribute [simps! coe] mgraph +attribute [simps! coe] AddMonoidHom.mgraph +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] coe_mgraph + +@[to_additive (attr := simp)] +lemma mem_mgraph {f : G →* H} {x : G × H} : x ∈ f.mgraph ↔ f x.1 = x.2 := .rfl + +@[to_additive] +lemma mgraph_eq_mrange_prod (f : G →* H) : f.mgraph = mrange ((id _).prod f) := by aesop + +/-- **Vertical line test** for monoid homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`. -/ +@[to_additive "**Vertical line test** for monoid homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some monoid homomorphism `f' : H → I`."] +lemma exists_mrange_eq_mgraph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →* I, mrange f = f'.mgraph := by + obtain ⟨f', hf'⟩ := exists_range_eq_graphOn_univ hf₁ hf + simp only [Set.ext_iff, Set.mem_range, mem_graphOn, mem_univ, true_and, Prod.forall] at hf' + use + { toFun := f' + map_one' := (hf' _ _).1 ⟨1, map_one _⟩ + map_mul' := by + simp_rw [hf₁.forall] + rintro g₁ g₂ + exact (hf' _ _).1 ⟨g₁ * g₂, by simp [Prod.ext_iff, (hf' (f _).1 _).1 ⟨_, rfl⟩]⟩ } + simpa [SetLike.ext_iff] using hf' + +/-- **Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some monoid +isomorphism `f' : H ≃ I`. -/ +@[to_additive "**Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` and every +\"horizontal line\" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +monoid isomorphism `f' : H ≃ I`."] +lemma exists_mulEquiv_mrange_eq_mgraph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃* I, mrange f = e.toMonoidHom.mgraph := by + obtain ⟨e₁, he₁⟩ := f.exists_mrange_eq_mgraph hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := (MulEquiv.prodComm.toMonoidHom.comp f).exists_mrange_eq_mgraph (by simpa) <| + by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + rw [SetLike.ext_iff] at he₁ he₂ + aesop (add simp [Prod.swap_eq_iff_eq_swap]) + exact ⟨ + { toFun := e₁ + map_mul' := e₁.map_mul' + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +end MonoidHom + +/-- **Vertical line test** for monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some monoid homomorphism `f : H → I`. -/ +@[to_additive "**Vertical line test** for additive monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` surjects onto the first +factor and `G` intersects every \"vertical line\" `{(h, i) | i : I}` at most once. Then `G` is the +graph of some monoid homomorphism `f : H → I`."] +lemma Submonoid.exists_eq_mgraph {G : Submonoid (H × I)} (hG₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →* I, G = f.mgraph := by + simpa using MonoidHom.exists_mrange_eq_mgraph hG₁.surjective + fun a b h ↦ congr_arg (Prod.snd ∘ G.subtype) (hG₁.injective h) + +/-- **Goursat's lemma** for monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃* I`. -/ +@[to_additive "**Goursat's lemma** for additive monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of additive monoids. Assume that the natural maps from +`G` to both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃+ I`."] +lemma Submonoid.exists_mulEquiv_eq_mgraph {G : Submonoid (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃* I, G = e.toMonoidHom.mgraph := by + simpa using MonoidHom.exists_mulEquiv_mrange_eq_mgraph hG₁.surjective hG₂.surjective + fun _ _ ↦ hG₁.injective.eq_iff.trans hG₂.injective.eq_iff.symm + +end Monoid + +section Group +variable [Group G] [Group H] [Group I] + +namespace MonoidHom + +/-- The graph of a group homomorphism as a subgroup. + +See also `MonoidHom.mgraph` for the graph as a submonoid. -/ +@[to_additive +"The graph of a group homomorphism as a subgroup. + +See also `AddMonoidHom.mgraph` for the graph as a submonoid."] +def graph (f : G →* H) : Subgroup (G × H) where + toSubmonoid := f.mgraph + inv_mem' {x} := by simp +contextual + +-- TODO: Can `to_additive` be smarter about `simps`? +attribute [simps! coe toSubmonoid] graph +attribute [simps! coe toAddSubmonoid] AddMonoidHom.graph +set_option linter.existingAttributeWarning false in +attribute [to_additive existing] coe_graph graph_toSubmonoid + +@[to_additive] +lemma mem_graph {f : G →* H} {x : G × H} : x ∈ f.mgraph ↔ f x.1 = x.2 := .rfl + +@[to_additive] +lemma graph_eq_range_prod (f : G →* H) : f.graph = range ((id _).prod f) := by aesop + +/-- **Vertical line test** for group homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some group homomorphism `f' : H → I`. -/ +@[to_additive "**Vertical line test** for group homomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on the +first factor and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` at most +once. Then the image of `f` is the graph of some group homomorphism `f' : H → I`."] +lemma exists_range_eq_graph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →* I, range f = f'.graph := by + simpa [SetLike.ext_iff] using exists_mrange_eq_mgraph hf₁ hf + +/-- **Line test** for group isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some group +isomorphism `f' : H ≃ I`. -/ +@[to_additive "**Line test** for monoid isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of groups. Assume that `f` is surjective on both +factors and that the image of `f` intersects every \"vertical line\" `{(h, i) | i : I}` and every +\"horizontal line\" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +group isomorphism `f' : H ≃ I`."] +lemma exists_mulEquiv_range_eq_graph {f : G →* H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃* I, range f = e.toMonoidHom.graph := by + simpa [SetLike.ext_iff] using exists_mulEquiv_mrange_eq_mgraph hf₁ hf₂ hf + +end MonoidHom + +/-- **Vertical line test** for group homomorphisms. + +Let `G ≤ H × I` be a subgroup of a product of monoids. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some monoid homomorphism `f : H → I`. -/ +@[to_additive "**Vertical line test** for additive monoid homomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that `G` surjects onto the first +factor and `G` intersects every \"vertical line\" `{(h, i) | i : I}` at most once. Then `G` is the +graph of some monoid homomorphism `f : H → I`."] +lemma Subgroup.exists_eq_graph {G : Subgroup (H × I)} (hG₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →* I, G = f.graph := by + simpa [SetLike.ext_iff] using Submonoid.exists_eq_mgraph hG₁ + +/-- **Goursat's lemma** for monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of monoids. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃* I`. -/ +@[to_additive "**Goursat's lemma** for additive monoid isomorphisms. + +Let `G ≤ H × I` be a submonoid of a product of additive monoids. Assume that the natural maps from +`G` to both factors are bijective. Then `G` is the graph of some isomorphism `f : H ≃+ I`."] +lemma Subgroup.exists_mulEquiv_eq_graph {G : Subgroup (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃* I, G = e.toMonoidHom.graph := by + simpa [SetLike.ext_iff] using Submonoid.exists_mulEquiv_eq_mgraph hG₁ hG₂ + +end Group diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index 02781d10d0067..e6e9519ae6483 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -149,6 +149,8 @@ theorem mk_div_mk [Div G] [Div H] (x₁ x₂ : G) (y₁ y₂ : H) : theorem swap_div [Div G] [Div H] (a b : G × H) : (a / b).swap = a.swap / b.swap := rfl +@[to_additive] lemma div_def [Div M] [Div N] (a b : M × N) : a / b = (a.1 / b.1, a.2 / b.2) := rfl + @[to_additive] instance instSemigroup [Semigroup M] [Semigroup N] : Semigroup (M × N) := { mul_assoc := fun _ _ _ => mk.inj_iff.mpr ⟨mul_assoc _ _ _, mul_assoc _ _ _⟩ } diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index f42386ead9042..98f729aaec3e6 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -5,12 +5,13 @@ Authors: Johannes Hölzl -/ import Mathlib.Logic.Function.Defs import Mathlib.Logic.Function.Iterate +import Aesop import Mathlib.Tactic.Inhabit /-! # Extra facts about `Prod` -This file defines `Prod.swap : α × β → β × α` and proves various simple lemmas about `Prod`. +This file proves various simple lemmas about `Prod`. It also defines better delaborators for product projections. -/ @@ -20,6 +21,8 @@ variable {α : Type*} {β : Type*} {γ : Type*} {δ : Type*} namespace Prod +lemma swap_eq_iff_eq_swap {x : α × β} {y : β × α} : x.swap = y ↔ x = y.swap := by aesop + def mk.injArrow {x₁ : α} {y₁ : β} {x₂ : α} {y₂ : β} : (x₁, y₁) = (x₂, y₂) → ∀ ⦃P : Sort*⦄, (x₁ = x₂ → y₁ = y₂ → P) → P := fun h₁ _ h₂ ↦ Prod.noConfusion h₁ h₂ diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 0a7bacfba799b..77dbaaa172001 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -1665,4 +1665,55 @@ lemma bijOn_swap (ha : a ∈ s) (hb : b ∈ s) : BijOn (swap a b) s s := end Equiv +/-! ### Vertical line test -/ + +namespace Set + +/-- **Vertical line test** for functions. + +Let `f : α → β × γ` be a function to a product. Assume that `f` is surjective on the first factor +and that the image of `f` intersects every "vertical line" `{(b, c) | c : γ}` at most once. +Then the image of `f` is the graph of some monoid homomorphism `f' : β → γ`. -/ +lemma exists_range_eq_graphOn_univ {f : α → β × γ} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : β → γ, range f = univ.graphOn f' := by + refine ⟨fun h ↦ (f (hf₁ h).choose).snd, ?_⟩ + ext x + simp only [mem_range, comp_apply, mem_graphOn, mem_univ, true_and] + refine ⟨?_, fun hi ↦ ⟨(hf₁ x.1).choose, Prod.ext (hf₁ x.1).choose_spec hi⟩⟩ + rintro ⟨g, rfl⟩ + exact hf _ _ (hf₁ (f g).1).choose_spec + +/-- **Line test** for equivalences. + +Let `f : α → β × γ` be a homomorphism to a product of monoids. Assume that `f` is surjective on both +factors and that the image of `f` intersects every "vertical line" `{(b, c) | c : γ}` and every +"horizontal line" `{(b, c) | b : β}` at most once. Then the image of `f` is the graph of some +equivalence `f' : β ≃ γ`. -/ +lemma exists_equiv_range_eq_graphOn_univ {f : α → β × γ} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : β ≃ γ, range f = univ.graphOn e := by + obtain ⟨e₁, he₁⟩ := exists_range_eq_graphOn_univ hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := exists_range_eq_graphOn_univ (f := Equiv.prodComm _ _ ∘ f) (by simpa) <| + by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + rw [Set.ext_iff] at he₁ he₂ + aesop (add simp [Prod.swap_eq_iff_eq_swap]) + exact ⟨ + { toFun := e₁ + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +/-- **Vertical line test** for functions. + +Let `s : Set (β × γ)` be a set in a product. Assume that `s` maps bijectively to the first factor. +Then `s` is the graph of some function `f : β → γ`. -/ +lemma exists_eq_mgraphOn_univ {s : Set (β × γ)} + (hs₁ : Bijective (Prod.fst ∘ (Subtype.val : s → β × γ))) : ∃ f : β → γ, s = univ.graphOn f := by + simpa using exists_range_eq_graphOn_univ hs₁.surjective + fun a b h ↦ congr_arg (Prod.snd ∘ (Subtype.val : s → β × γ)) (hs₁.injective h) + +end Set + set_option linter.style.longFile 1800 diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 7bbb4ac2f4b65..23732c0a35284 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl, Mario Carneiro, Kevin Buzzard, Yury Kudryashov, Eric Wieser -/ import Mathlib.Algebra.Algebra.Prod +import Mathlib.Algebra.Group.Graph import Mathlib.LinearAlgebra.Span.Basic import Mathlib.Order.PartialSups @@ -965,3 +966,84 @@ theorem graph_eq_range_prod : f.graph = range (LinearMap.id.prod f) := by end Graph end LinearMap + +section LineTest + +open Set Function + +variable {R S G H I : Type*} + [Semiring R] [Semiring S] {σ : R →+* S} [RingHomSurjective σ] + [AddCommMonoid G] [Module R G] + [AddCommMonoid H] [Module S H] + [AddCommMonoid I] [Module S I] + +/-- **Vertical line test** for module homomorphisms. + +Let `f : G → H × I` be a linear (or semilinear) map to a product. Assume that `f` is surjective on +the first factor and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` at +most once. Then the image of `f` is the graph of some linear map `f' : H → I`. -/ +lemma LinearMap.exists_range_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : + ∃ f' : H →ₗ[S] I, LinearMap.range f = LinearMap.graph f' := by + obtain ⟨f', hf'⟩ := AddMonoidHom.exists_mrange_eq_mgraph (I := I) (f := f) hf₁ hf + simp only [SetLike.ext_iff, AddMonoidHom.mem_mrange, AddMonoidHom.coe_coe, + AddMonoidHom.mem_mgraph] at hf' + use + { toFun := f'.toFun + map_add' := f'.map_add' + map_smul' := by + intro s h + simp only [ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, RingHom.id_apply] + refine (hf' (s • h, _)).mp ?_ + rw [← Prod.smul_mk, ← LinearMap.mem_range] + apply Submodule.smul_mem + rw [LinearMap.mem_range, hf'] } + ext x + simpa only [mem_range, Eq.comm, ZeroHom.toFun_eq_coe, AddMonoidHom.toZeroHom_coe, mem_graph_iff, + coe_mk, AddHom.coe_mk, AddMonoidHom.coe_coe, Set.mem_range] using hf' x + +/-- **Vertical line test** for module homomorphisms. + +Let `G ≤ H × I` be a submodule of a product of modules. Assume that `G` maps bijectively to the +first factor. Then `G` is the graph of some module homomorphism `f : H →ₗ[R] I`. -/ +lemma Submodule.exists_eq_graph {G : Submodule S (H × I)} (hf₁ : Bijective (Prod.fst ∘ G.subtype)) : + ∃ f : H →ₗ[S] I, G = LinearMap.graph f := by + simpa only [range_subtype] using LinearMap.exists_range_eq_graph hf₁.surjective + (fun a b h ↦ congr_arg (Prod.snd ∘ G.subtype) (hf₁.injective h)) + +/-- **Line test** for module isomorphisms. + +Let `f : G → H × I` be a homomorphism to a product of modules. Assume that `f` is surjective onto +both factors and that the image of `f` intersects every "vertical line" `{(h, i) | i : I}` and every +"horizontal line" `{(h, i) | h : H}` at most once. Then the image of `f` is the graph of some +module isomorphism `f' : H ≃ I`. -/ +lemma LinearMap.exists_linearEquiv_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) + (hf₂ : Surjective (Prod.snd ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 ↔ (f g₁).2 = (f g₂).2) : + ∃ e : H ≃ₗ[S] I, range f = e.toLinearMap.graph := by + obtain ⟨e₁, he₁⟩ := f.exists_range_eq_graph hf₁ fun _ _ ↦ (hf _ _).1 + obtain ⟨e₂, he₂⟩ := ((LinearEquiv.prodComm _ _ _).toLinearMap.comp f).exists_range_eq_graph + (by simpa) <| by simp [hf] + have he₁₂ h i : e₁ h = i ↔ e₂ i = h := by + simp only [SetLike.ext_iff, LinearMap.mem_graph_iff] at he₁ he₂ + rw [Eq.comm, ← he₁ (h, i), Eq.comm, ← he₂ (i, h)] + simp only [mem_range, coe_comp, LinearEquiv.coe_coe, Function.comp_apply, + LinearEquiv.prodComm_apply, Prod.swap_eq_iff_eq_swap, Prod.swap_prod_mk] + exact ⟨ + { toFun := e₁ + map_smul' := e₁.map_smul' + map_add' := e₁.map_add' + invFun := e₂ + left_inv := fun h ↦ by rw [← he₁₂] + right_inv := fun i ↦ by rw [he₁₂] }, he₁⟩ + +/-- **Goursat's lemma** for module isomorphisms. + +Let `G ≤ H × I` be a submodule of a product of modules. Assume that the natural maps from `G` to +both factors are bijective. Then `G` is the graph of some module isomorphism `f : H ≃ I`. -/ +lemma Submodule.exists_equiv_eq_graph {G : Submodule S (H × I)} + (hG₁ : Bijective (Prod.fst ∘ G.subtype)) (hG₂ : Bijective (Prod.snd ∘ G.subtype)) : + ∃ e : H ≃ₗ[S] I, G = e.toLinearMap.graph := by + simpa only [range_subtype] using LinearMap.exists_linearEquiv_eq_graph + hG₁.surjective hG₂.surjective fun _ _ ↦ hG₁.injective.eq_iff.trans hG₂.injective.eq_iff.symm + +end LineTest From 04cd3d0afc11a0bcb07934a8e2f12a72793b2f28 Mon Sep 17 00:00:00 2001 From: damiano Date: Mon, 25 Nov 2024 15:32:01 +0000 Subject: [PATCH 149/250] fix: remove `ref: master` in zulip emoji reaction script (#19465) This way, the script will extract the python file from the current PR, allowing testing and debugging of changes in its behaviour. --- .github/workflows/maintainer_bors.yml | 1 - 1 file changed, 1 deletion(-) diff --git a/.github/workflows/maintainer_bors.yml b/.github/workflows/maintainer_bors.yml index 67e3596b6a67e..10a12e5f3a439 100644 --- a/.github/workflows/maintainer_bors.yml +++ b/.github/workflows/maintainer_bors.yml @@ -103,7 +103,6 @@ jobs: - uses: actions/checkout@v4 with: - ref: master sparse-checkout: | scripts/zulip_emoji_merge_delegate.py From 15ea3626fdf7430332c8b5a605e53498a0094512 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Mon, 25 Nov 2024 16:00:55 +0000 Subject: [PATCH 150/250] chore(Topology/../LiminfLimsup): change assumptions of 2 instances (#18919) - Assume `LinearOrder` instead of `Preorder`. - Assume `ClosedIciTopology`/`ClosedIicTopology` instead of `OrderTopology`. `OrderTopology` is usually not the right topology for non-linear orders, so restricting from `Preorder` to `LinearOrder` doesn't break anything. --- .../Topology/Algebra/Order/LiminfLimsup.lean | 20 +++++++++---------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean index bc4753090a8e3..9f2b202d56789 100644 --- a/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean +++ b/Mathlib/Topology/Algebra/Order/LiminfLimsup.lean @@ -130,20 +130,18 @@ instance (priority := 100) OrderTop.to_BoundedLENhdsClass [OrderTop α] : Bounde instance (priority := 100) OrderBot.to_BoundedGENhdsClass [OrderBot α] : BoundedGENhdsClass α := ⟨fun _a ↦ isBounded_ge_of_bot⟩ --- See note [lower instance priority] -instance (priority := 100) OrderTopology.to_BoundedLENhdsClass [IsDirected α (· ≤ ·)] - [OrderTopology α] : BoundedLENhdsClass α := - ⟨fun a ↦ - ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| - Exists.imp fun _b ↦ ge_mem_nhds⟩ +end Preorder -- See note [lower instance priority] -instance (priority := 100) OrderTopology.to_BoundedGENhdsClass [IsDirected α (· ≥ ·)] - [OrderTopology α] : BoundedGENhdsClass α := - ⟨fun a ↦ ((isBot_or_exists_lt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| - Exists.imp fun _b ↦ le_mem_nhds⟩ +instance (priority := 100) BoundedLENhdsClass.of_closedIciTopology [LinearOrder α] + [TopologicalSpace α] [ClosedIciTopology α] : BoundedLENhdsClass α := + ⟨fun a ↦ ((isTop_or_exists_gt a).elim fun h ↦ ⟨a, Eventually.of_forall h⟩) <| + Exists.imp fun _b ↦ eventually_le_nhds⟩ -end Preorder +-- See note [lower instance priority] +instance (priority := 100) BoundedGENhdsClass.of_closedIicTopology [LinearOrder α] + [TopologicalSpace α] [ClosedIicTopology α] : BoundedGENhdsClass α := + inferInstanceAs <| BoundedGENhdsClass αᵒᵈᵒᵈ section LiminfLimsup From 435fc2bcedc84fd97664f7c4ece07d979765063c Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Mon, 25 Nov 2024 16:18:55 +0000 Subject: [PATCH 151/250] feat(AlgebraicGeometry): covers of schemes over a base (#19096) We introduce a class on covers containing the data of structure morphisms over a base and provide basic constructions. Co-authored-by: Christian Merten <136261474+chrisflav@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/AlgebraicGeometry/Cover/Over.lean | 183 ++++++++++++++++++ .../Morphisms/UnderlyingMap.lean | 12 ++ Mathlib/AlgebraicGeometry/Over.lean | 7 + Mathlib/CategoryTheory/Comma/OverClass.lean | 15 ++ .../Limits/MorphismProperty.lean | 2 +- .../MorphismProperty/Comma.lean | 8 + 7 files changed, 227 insertions(+), 1 deletion(-) create mode 100644 Mathlib/AlgebraicGeometry/Cover/Over.lean diff --git a/Mathlib.lean b/Mathlib.lean index 979ce54ccf805..c6a763cc2c95e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -931,6 +931,7 @@ import Mathlib.AlgebraicGeometry.AffineScheme import Mathlib.AlgebraicGeometry.AffineSpace import Mathlib.AlgebraicGeometry.Cover.MorphismProperty import Mathlib.AlgebraicGeometry.Cover.Open +import Mathlib.AlgebraicGeometry.Cover.Over import Mathlib.AlgebraicGeometry.EllipticCurve.Affine import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Basic import Mathlib.AlgebraicGeometry.EllipticCurve.DivisionPolynomial.Degree diff --git a/Mathlib/AlgebraicGeometry/Cover/Over.lean b/Mathlib/AlgebraicGeometry/Cover/Over.lean new file mode 100644 index 0000000000000..e6b2a8b4e0e23 --- /dev/null +++ b/Mathlib/AlgebraicGeometry/Cover/Over.lean @@ -0,0 +1,183 @@ +/- +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.UnderlyingMap +import Mathlib.CategoryTheory.Limits.MorphismProperty + +/-! + +# Covers of schemes over a base + +In this file we define the typeclass `Cover.Over`. For a cover `𝒰` of an `S`-scheme `X`, +the datum `𝒰.Over S` contains `S`-scheme structures on the components of `𝒰` and asserts +that the component maps are morphisms of `S`-schemes. + +We provide instances of `𝒰.Over S` for standard constructions on covers. + +-/ + +universe v u + +noncomputable section + +open CategoryTheory Limits + +namespace AlgebraicGeometry.Scheme + +variable {P : MorphismProperty Scheme.{u}} (S : Scheme.{u}) + +/-- Bundle an `S`-scheme with `P` into an object of `P.Over ⊤ S`. -/ +abbrev asOverProp (X : Scheme.{u}) (S : Scheme.{u}) [X.Over S] (h : P (X ↘ S)) : P.Over ⊤ S := + ⟨X.asOver S, h⟩ + +/-- Bundle an `S`-morphism of `S`-scheme with `P` into a morphism in `P.Over ⊤ S`. -/ +abbrev Hom.asOverProp {X Y : Scheme.{u}} (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] + [f.IsOver S] {hX : P (X ↘ S)} {hY : P (Y ↘ S)} : X.asOverProp S hX ⟶ Y.asOverProp S hY := + ⟨f.asOver S, trivial, trivial⟩ + +/-- A `P`-cover of a scheme `X` over `S` is a cover, where the components are over `S` and the +component maps commute with the structure morphisms. -/ +protected class Cover.Over {P : MorphismProperty Scheme.{u}} {X : Scheme.{u}} [X.Over S] + (𝒰 : X.Cover P) where + over (j : 𝒰.J) : (𝒰.obj j).Over S := by infer_instance + isOver_map (j : 𝒰.J) : (𝒰.map j).IsOver S := by infer_instance + +attribute [instance] Cover.Over.over Cover.Over.isOver_map + +instance [P.ContainsIdentities] [P.RespectsIso] {X Y : Scheme.{u}} (f : X ⟶ Y) [X.Over S] [Y.Over S] + [f.IsOver S] [IsIso f] : (coverOfIsIso (P := P) f).Over S where + over _ := inferInstanceAs <| X.Over S + isOver_map _ := inferInstanceAs <| f.IsOver S + +section + +variable [P.IsStableUnderBaseChange] [IsJointlySurjectivePreserving P] +variable {X W : Scheme.{u}} (𝒰 : X.Cover P) (f : W ⟶ X) [W.Over S] [X.Over S] + [𝒰.Over S] [f.IsOver S] + +/-- The pullback of a cover of `S`-schemes along a morphism of `S`-schemes. This is not +definitionally equal to `AlgebraicGeometry.Scheme.Cover.pullbackCover`, as here we take +the pullback in `Over S`, whose underlying scheme is only isomorphic but not equal to the +pullback in `Scheme`. -/ +@[simps] +def Cover.pullbackCoverOver : W.Cover P where + J := 𝒰.J + obj x := (pullback (f.asOver S) ((𝒰.map x).asOver S)).left + map x := (pullback.fst (f.asOver S) ((𝒰.map x).asOver S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (Over.forget S) (f.asOver S) ((𝒰.map _).asOver S)).inv) + (PreservesPullback.iso_inv_fst _ _ _) x).mp ((𝒰.pullbackCover f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, ← PreservesPullback.iso_hom_fst, P.cancel_left_of_respectsIso] + exact P.pullback_fst _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOver S f).obj j).Over S where + hom := (pullback (f.asOver S) ((𝒰.map j).asOver S)).hom + +instance : (𝒰.pullbackCoverOver S f).Over S where + isOver_map j := { comp_over := by exact Over.w (pullback.fst (f.asOver S) ((𝒰.map j).asOver S)) } + +/-- A variant of `AlgebraicGeometry.Scheme.Cover.pullbackCoverOver` with the arguments in the +fiber products flipped. -/ +@[simps] +def Cover.pullbackCoverOver' : W.Cover P where + J := 𝒰.J + obj x := (pullback ((𝒰.map x).asOver S) (f.asOver S)).left + map x := (pullback.snd ((𝒰.map x).asOver S) (f.asOver S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover' f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (Over.forget S) ((𝒰.map _).asOver S) (f.asOver S)).inv) + (PreservesPullback.iso_inv_snd _ _ _) x).mp ((𝒰.pullbackCover' f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, ← PreservesPullback.iso_hom_snd, P.cancel_left_of_respectsIso] + exact P.pullback_snd _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOver' S f).obj j).Over S where + hom := (pullback ((𝒰.map j).asOver S) (f.asOver S)).hom + +instance : (𝒰.pullbackCoverOver' S f).Over S where + isOver_map j := { comp_over := by exact Over.w (pullback.snd ((𝒰.map j).asOver S) (f.asOver S)) } + +variable {Q : MorphismProperty Scheme.{u}} [Q.HasOfPostcompProperty Q] + [Q.IsStableUnderBaseChange] [Q.IsStableUnderComposition] + +variable (hX : Q (X ↘ S)) (hW : Q (W ↘ S)) (hQ : ∀ j, Q (𝒰.obj j ↘ S)) + +/-- The pullback of a cover of `S`-schemes with `Q` along a morphism of `S`-schemes. This is not +definitionally equal to `AlgebraicGeometry.Scheme.Cover.pullbackCover`, as here we take +the pullback in `Q.Over ⊤ S`, whose underlying scheme is only isomorphic but not equal to the +pullback in `Scheme`. -/ +@[simps (config := .lemmasOnly)] +def Cover.pullbackCoverOverProp : W.Cover P where + J := 𝒰.J + obj x := (pullback (f.asOverProp (hX := hW) (hY := hX) S) + ((𝒰.map x).asOverProp (hX := hQ x) (hY := hX) S)).left + map x := (pullback.fst (f.asOverProp S) ((𝒰.map x).asOverProp S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (MorphismProperty.Over.forget Q _ _ ⋙ Over.forget S) + (f.asOverProp S) ((𝒰.map _).asOverProp S)).inv) + (PreservesPullback.iso_inv_fst _ _ _) x).mp ((𝒰.pullbackCover f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, MorphismProperty.Comma.toCommaMorphism_eq_hom, + ← MorphismProperty.Comma.forget_map, ← Functor.comp_map] + rw [← PreservesPullback.iso_hom_fst, P.cancel_left_of_respectsIso] + exact P.pullback_fst _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOverProp S f hX hW hQ).obj j).Over S where + hom := (pullback (f.asOverProp (hX := hW) (hY := hX) S) + ((𝒰.map j).asOverProp (hX := hQ j) (hY := hX) S)).hom + +instance : (𝒰.pullbackCoverOverProp S f hX hW hQ).Over S where + isOver_map j := + { comp_over := by exact (pullback.fst (f.asOverProp S) ((𝒰.map j).asOverProp S)).w } + +/-- A variant of `AlgebraicGeometry.Scheme.Cover.pullbackCoverOverProp` with the arguments in the +fiber products flipped. -/ +@[simps (config := .lemmasOnly)] +def Cover.pullbackCoverOverProp' : W.Cover P where + J := 𝒰.J + obj x := (pullback ((𝒰.map x).asOverProp (hX := hQ x) (hY := hX) S) + (f.asOverProp (hX := hW) (hY := hX) S)).left + map x := (pullback.snd ((𝒰.map x).asOverProp S) (f.asOverProp S)).left + f x := 𝒰.f (f.base x) + covers x := (mem_range_iff_of_surjective ((𝒰.pullbackCover' f).map (𝒰.f (f.base x))) _ + ((PreservesPullback.iso (MorphismProperty.Over.forget Q _ _ ⋙ Over.forget S) + ((𝒰.map _).asOverProp S) (f.asOverProp S)).inv) + (PreservesPullback.iso_inv_snd _ _ _) x).mp ((𝒰.pullbackCover' f).covers x) + map_prop j := by + dsimp only + rw [← Over.forget_map, MorphismProperty.Comma.toCommaMorphism_eq_hom, + ← MorphismProperty.Comma.forget_map, ← Functor.comp_map] + rw [← PreservesPullback.iso_hom_snd, P.cancel_left_of_respectsIso] + exact P.pullback_snd _ _ (𝒰.map_prop j) + +instance (j : 𝒰.J) : ((𝒰.pullbackCoverOverProp' S f hX hW hQ).obj j).Over S where + hom := (pullback ((𝒰.map j).asOverProp (hX := hQ j) (hY := hX) S) + (f.asOverProp (hX := hW) (hY := hX) S)).hom + +instance : (𝒰.pullbackCoverOverProp' S f hX hW hQ).Over S where + isOver_map j := + { comp_over := by exact (pullback.snd ((𝒰.map j).asOverProp S) (f.asOverProp S)).w } + +end + +variable [P.IsStableUnderComposition] +variable {X : Scheme.{u}} (𝒰 : X.Cover P) (𝒱 : ∀ x, (𝒰.obj x).Cover P) + [X.Over S] [𝒰.Over S] [∀ x, (𝒱 x).Over S] + +instance (j : (𝒰.bind 𝒱).J) : ((𝒰.bind 𝒱).obj j).Over S := + inferInstanceAs <| ((𝒱 j.1).obj j.2).Over S + +instance {X : Scheme.{u}} (𝒰 : X.Cover P) (𝒱 : ∀ x, (𝒰.obj x).Cover P) + [X.Over S] [𝒰.Over S] [∀ x, (𝒱 x).Over S] : (𝒰.bind 𝒱).Over S where + over := fun ⟨i, j⟩ ↦ inferInstanceAs <| ((𝒱 i).obj j).Over S + isOver_map := fun ⟨i, j⟩ ↦ { comp_over := by simp } + +end AlgebraicGeometry.Scheme diff --git a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean index dbcc339026d05..065d25bd88b11 100644 --- a/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean +++ b/Mathlib/AlgebraicGeometry/Morphisms/UnderlyingMap.lean @@ -83,6 +83,18 @@ instance surjective_isLocalAtTarget : IsLocalAtTarget @Surjective := by obtain ⟨⟨y, _⟩, hy⟩ := hf i ⟨x, hxi⟩ exact ⟨y, congr(($hy).1)⟩ +@[simp] +lemma range_eq_univ [Surjective f] : Set.range f.base = Set.univ := by + simpa [Set.range_eq_univ] using f.surjective + +lemma range_eq_range_of_surjective {S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) (e : X ⟶ Y) + [Surjective e] (hge : e ≫ g = f) : Set.range f.base = Set.range g.base := by + rw [← hge] + simp [Set.range_comp] + +lemma mem_range_iff_of_surjective {S : Scheme.{u}} (f : X ⟶ S) (g : Y ⟶ S) (e : X ⟶ Y) + [Surjective e] (hge : e ≫ g = f) (s : S) : s ∈ Set.range f.base ↔ s ∈ Set.range g.base := by + rw [range_eq_range_of_surjective f g e hge] end Surjective section Injective diff --git a/Mathlib/AlgebraicGeometry/Over.lean b/Mathlib/AlgebraicGeometry/Over.lean index b0f3bcaef6d0e..15349f1b13a3e 100644 --- a/Mathlib/AlgebraicGeometry/Over.lean +++ b/Mathlib/AlgebraicGeometry/Over.lean @@ -47,4 +47,11 @@ lemma Hom.isOver_iff [X.Over S] [Y.Over S] {f : X ⟶ Y} : f.IsOver S ↔ f ≫ /-! Also note the existence of `CategoryTheory.IsOverTower X Y S`. -/ +/-- Given `X.Over S`, this is the bundled object of `Over S`. -/ +abbrev asOver (X S : Scheme.{u}) [X.Over S] := OverClass.asOver X S + +/-- Given a morphism `X ⟶ Y` with `f.IsOver S`, this is the bundled morphism in `Over S`. -/ +abbrev Hom.asOver (f : X.Hom Y) (S : Scheme.{u}) [X.Over S] [Y.Over S] [f.IsOver S] := + OverClass.asOverHom S f + end AlgebraicGeometry.Scheme diff --git a/Mathlib/CategoryTheory/Comma/OverClass.lean b/Mathlib/CategoryTheory/Comma/OverClass.lean index c9ad20a1fd78e..a6b1fc35caeb2 100644 --- a/Mathlib/CategoryTheory/Comma/OverClass.lean +++ b/Mathlib/CategoryTheory/Comma/OverClass.lean @@ -120,7 +120,22 @@ instance [OverClass X S] [IsOverTower X S S'] [IsOverTower Y S S'] [HomIsOver f S] : HomIsOver f S' := homIsOver_of_isOverTower f S S' +variable (X) in /-- Bundle `X` with an `OverClass X S` instance into `Over S`. -/ +@[simps! hom left] def OverClass.asOver [OverClass X S] : Over S := Over.mk (X ↘ S) +/-- Bundle a morphism `f : X ⟶ Y` with `HomIsOver f S` into a morphism in `Over S`. -/ +@[simps! left] +def OverClass.asOverHom [OverClass X S] [OverClass Y S] (f : X ⟶ Y) [HomIsOver f S] : + OverClass.asOver X S ⟶ OverClass.asOver Y S := + Over.homMk f (comp_over f S) + +@[simps] +instance OverClass.fromOver {S : C} (X : Over S) : OverClass X.left S where + hom := X.hom + +instance {S : C} {X Y : Over S} (f : X ⟶ Y) : HomIsOver f.left S where + comp_over := Over.w f + end CategoryTheory diff --git a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean index 8f764982cf362..b4bb842c72dd2 100644 --- a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean +++ b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean @@ -134,7 +134,7 @@ instance [P.ContainsIdentities] : HasTerminal (P.Over ⊤ X) := /-- If `P` is stable under composition, base change and satisfies post-cancellation, `Over.forget P ⊤ X` creates pullbacks. -/ -noncomputable def createsLimitsOfShape_walkingCospan [HasPullbacks T] +noncomputable instance createsLimitsOfShape_walkingCospan [HasPullbacks T] [P.IsStableUnderComposition] [P.IsStableUnderBaseChange] [P.HasOfPostcompProperty P] : CreatesLimitsOfShape WalkingCospan (Over.forget P ⊤ X) := haveI : HasLimitsOfShape WalkingCospan (Comma (𝟭 T) (Functor.fromPUnit X)) := diff --git a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean index 0662c8046f120..6b093565af86f 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean @@ -117,6 +117,8 @@ instance : Category (P.Comma L R Q W) where id X := X.id comp f g := f.comp g +lemma toCommaMorphism_eq_hom {X Y : P.Comma L R Q W} (f : X ⟶ Y) : f.toCommaMorphism = f.hom := rfl + /-- Alternative `ext` lemma for `Comma.Hom`. -/ @[ext] lemma Hom.ext' {X Y : P.Comma L R Q W} {f g : X ⟶ Y} (h : f.hom = g.hom) : @@ -220,6 +222,9 @@ protected abbrev Over : Type _ := protected abbrev Over.forget : P.Over Q X ⥤ Over X := Comma.forget (Functor.id T) (Functor.fromPUnit.{0} X) P Q ⊤ +instance : (Over.forget P ⊤ X).Faithful := inferInstanceAs <| (Comma.forget _ _ _ _ _).Faithful +instance : (Over.forget P ⊤ X).Full := inferInstanceAs <| (Comma.forget _ _ _ _ _).Full + variable {P Q X} /-- Construct a morphism in `P.Over Q X` from a morphism in `Over.X`. -/ @@ -261,6 +266,9 @@ protected abbrev Under : Type _ := protected abbrev Under.forget : P.Under Q X ⥤ Under X := Comma.forget (Functor.fromPUnit.{0} X) (Functor.id T) P ⊤ Q +instance : (Under.forget P ⊤ X).Faithful := inferInstanceAs <| (Comma.forget _ _ _ _ _).Faithful +instance : (Under.forget P ⊤ X).Full := inferInstanceAs <| (Comma.forget _ _ _ _ _).Full + variable {P Q X} /-- Construct a morphism in `P.Under Q X` from a morphism in `Under.X`. -/ From 551481f4b8470d7f67ced27d202a86942a642a49 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Mon, 25 Nov 2024 16:30:32 +0000 Subject: [PATCH 152/250] feat: ring with discrete PrimeSpectrum (#18891) Prove two characterizations of commutative semirings R whose prime spectra have discrete topology: + R has finitely many prime ideals, and every prime ideal is maximal; + R has finitely many maximal ideals, the intersection of which (= Jacobson radical) is contained in (and therefore equal to) the nilradical. From the first characterization it is easy to show the spectrum of an Artinian ring has discrete topology, but Artinian doesn't currently import PrimeSpectrum. Also show that if R is a commutative ring and Spec R has discrete topology, then the canonical RingHom from R to the product of localizations at maximal ideals (which is always injective) is an isomorphism. This is not a characterization, because it is satisfied by every local ring. (But it's a characterization if we replace "maximal" by "prime".) For this purpose we prove that, if a basic open D(f) in Spec R is a singleton {p}, then localization away from f is the same as localization at p. A key fact used is that if an ideal in a commutative semiring is disjoint from a submonoid, then there exists a maximal such ideal containing the ideal, and any maximal such ideal is prime. Also upgrade `exists_idempotent_basicOpen_eq_of_isClopen` to `existsUnique`. --- .../Algebra/Group/Submonoid/Membership.lean | 16 ++ Mathlib/Algebra/Ring/Idempotents.lean | 11 + .../PrimeSpectrum/Basic.lean | 196 ++++++++++++++++-- Mathlib/RingTheory/Artinian.lean | 7 +- Mathlib/RingTheory/Ideal/Maximal.lean | 20 ++ Mathlib/RingTheory/Ideal/Span.lean | 19 ++ Mathlib/RingTheory/Idempotents.lean | 2 + Mathlib/RingTheory/Localization/AtPrime.lean | 20 ++ .../RingTheory/Localization/Away/Basic.lean | 57 ++--- .../RingTheory/Localization/Away/Lemmas.lean | 10 +- Mathlib/RingTheory/Localization/Basic.lean | 12 +- Mathlib/RingTheory/Localization/Defs.lean | 5 + Mathlib/RingTheory/Nilpotent/Basic.lean | 8 + Mathlib/RingTheory/PrimeSpectrum.lean | 8 + Mathlib/Topology/Order.lean | 5 + 15 files changed, 337 insertions(+), 59 deletions(-) diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index 2f07619bf80ce..b48b123c3e191 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -8,6 +8,7 @@ import Mathlib.Algebra.FreeMonoid.Basic import Mathlib.Algebra.Group.Submonoid.MulOpposite import Mathlib.Algebra.Group.Submonoid.Operations import Mathlib.Algebra.GroupWithZero.Divisibility +import Mathlib.Algebra.Ring.Idempotents import Mathlib.Algebra.Ring.Int.Defs import Mathlib.Data.Finset.NoncommProd import Mathlib.Data.Nat.Cast.Basic @@ -420,6 +421,21 @@ lemma powers_le {n : M} {P : Submonoid M} : powers n ≤ P ↔ n ∈ P := by sim lemma powers_one : powers (1 : M) = ⊥ := bot_unique <| powers_le.2 <| one_mem _ +theorem _root_.IsIdempotentElem.coe_powers {a : M} (ha : IsIdempotentElem a) : + (Submonoid.powers a : Set M) = {1, a} := + let S : Submonoid M := + { carrier := {1, a}, + mul_mem' := by + rintro _ _ (rfl|rfl) (rfl|rfl) + · rw [one_mul]; exact .inl rfl + · rw [one_mul]; exact .inr rfl + · rw [mul_one]; exact .inr rfl + · rw [ha]; exact .inr rfl + one_mem' := .inl rfl } + suffices Submonoid.powers a = S from congr_arg _ this + le_antisymm (Submonoid.powers_le.mpr <| .inr rfl) + (by rintro _ (rfl|rfl); exacts [one_mem _, Submonoid.mem_powers _]) + /-- The submonoid generated by an element is a group if that element has finite order. -/ abbrev groupPowers {x : M} {n : ℕ} (hpos : 0 < n) (hx : x ^ n = 1) : Group (powers x) where inv x := x ^ (n - 1) diff --git a/Mathlib/Algebra/Ring/Idempotents.lean b/Mathlib/Algebra/Ring/Idempotents.lean index 80a2c2e0be568..8b7ce4e91aadd 100644 --- a/Mathlib/Algebra/Ring/Idempotents.lean +++ b/Mathlib/Algebra/Ring/Idempotents.lean @@ -67,6 +67,17 @@ theorem one_sub {p : R} (h : IsIdempotentElem p) : IsIdempotentElem (1 - p) := b theorem one_sub_iff {p : R} : IsIdempotentElem (1 - p) ↔ IsIdempotentElem p := ⟨fun h => sub_sub_cancel 1 p ▸ h.one_sub, IsIdempotentElem.one_sub⟩ +theorem add_sub_mul_of_commute {R} [Ring R] {p q : R} (h : Commute p q) + (hp : IsIdempotentElem p) (hq : IsIdempotentElem q) : + IsIdempotentElem (p + q - p * q) := by + convert (hp.one_sub.mul_of_commute ?_ hq.one_sub).one_sub using 1 + · simp_rw [sub_mul, mul_sub, one_mul, mul_one, sub_sub, sub_sub_cancel, add_sub, add_comm] + · simp_rw [commute_iff_eq, sub_mul, mul_sub, one_mul, mul_one, sub_sub, add_sub, add_comm, h.eq] + +theorem add_sub_mul {R} [CommRing R] {p q : R} (hp : IsIdempotentElem p) (hq : IsIdempotentElem q) : + IsIdempotentElem (p + q - p * q) := + add_sub_mul_of_commute (mul_comm p q) hp hq + theorem pow {p : N} (n : ℕ) (h : IsIdempotentElem p) : IsIdempotentElem (p ^ n) := Nat.recOn n ((pow_zero p).symm ▸ one) fun n _ => show p ^ n.succ * p ^ n.succ = p ^ n.succ by diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 8121592d7b482..6a4899d2598b7 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.Localization.Away.Basic +import Mathlib.RingTheory.Localization.Away.Lemmas import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -212,6 +212,31 @@ instance compactSpace : CompactSpace (PrimeSpectrum R) := by simp_rw [hI, ← zeroLocus_iSup, zeroLocus_empty_iff_eq_top, ← top_le_iff] at S_empty ⊢ exact Ideal.isCompactElement_top.exists_finset_of_le_iSup _ _ S_empty +/-- The prime spectrum of a semiring has discrete Zariski topology iff it is finite and +all primes are maximal. -/ +theorem discreteTopology_iff_finite_and_isPrime_imp_isMaximal : DiscreteTopology (PrimeSpectrum R) ↔ + Finite (PrimeSpectrum R) ∧ ∀ I : Ideal R, I.IsPrime → I.IsMaximal := + ⟨fun _ ↦ ⟨finite_of_compact_of_discrete, fun I hI ↦ (isClosed_singleton_iff_isMaximal ⟨I, hI⟩).mp + <| discreteTopology_iff_forall_isClosed.mp ‹_› _⟩, fun ⟨_, h⟩ ↦ .of_finite_of_isClosed_singleton + fun p ↦ (isClosed_singleton_iff_isMaximal p).mpr <| h _ p.2⟩ + +/-- The prime spectrum of a semiring has discrete Zariski topology iff there are only +finitely many maximal ideals and their intersection is contained in the nilradical. -/ +theorem discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical : + letI s := {I : Ideal R | I.IsMaximal} + DiscreteTopology (PrimeSpectrum R) ↔ Finite s ∧ sInf s ≤ nilradical R := + discreteTopology_iff_finite_and_isPrime_imp_isMaximal.trans <| by + rw [(equivSubtype R).finite_iff, ← Set.coe_setOf, Set.finite_coe_iff, Set.finite_coe_iff] + refine ⟨fun h ↦ ⟨h.1.subset fun _ h ↦ h.isPrime, nilradical_eq_sInf R ▸ sInf_le_sInf h.2⟩, + fun ⟨fin, le⟩ ↦ ?_⟩ + have hpm (I : Ideal R) (hI : I.IsPrime): I.IsMaximal := by + replace le := le.trans (nilradical_le_prime I) + rw [← fin.coe_toFinset, ← Finset.inf_id_eq_sInf, hI.inf_le'] at le + have ⟨M, hM, hMI⟩ := le + rw [fin.mem_toFinset] at hM + rwa [← hM.eq_of_le hI.1 hMI] + exact ⟨fin.subset hpm, hpm⟩ + section Comap variable {S' : Type*} [CommSemiring S'] @@ -501,6 +526,22 @@ lemma iSup_basicOpen_eq_top_iff' {s : Set R} : conv_rhs => rw [← Subtype.range_val (s := s), ← iSup_basicOpen_eq_top_iff] simp +theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] + {f : R} {p : PrimeSpectrum R} (h : (basicOpen f).1 = {p}) : + IsLocalization.Away f S ↔ IsLocalization.AtPrime S p.1 := + have : IsLocalization.AtPrime (Localization.Away f) p.1 := by + refine .of_le_of_exists_dvd (Submonoid.powers f) _ + (Submonoid.powers_le.mpr <| by apply h ▸ Set.mem_singleton p) fun r hr ↦ ?_ + contrapose! hr + simp_rw [← Ideal.mem_span_singleton] at hr + have ⟨q, prime, le, disj⟩ := Ideal.exists_le_prime_disjoint (Ideal.span {r}) + (Submonoid.powers f) (Set.disjoint_right.mpr hr) + have : ⟨q, prime⟩ ∈ (basicOpen f).1 := Set.disjoint_right.mp disj (Submonoid.mem_powers f) + rw [h, Set.mem_singleton_iff] at this + rw [← this] + exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) + IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) + end BasicOpen section Order @@ -669,14 +710,30 @@ lemma isCompact_isOpen_iff {s : Set (PrimeSpectrum R)} : fun ⟨s, e⟩ ↦ ⟨s, s.finite_toSet, by simpa using e.symm⟩⟩ lemma isCompact_isOpen_iff_ideal {s : Set (PrimeSpectrum R)} : - IsCompact s ∧ IsOpen s ↔ - ∃ I : Ideal R, I.FG ∧ (PrimeSpectrum.zeroLocus (I : Set R))ᶜ = s := by + IsCompact s ∧ IsOpen s ↔ ∃ I : Ideal R, I.FG ∧ (zeroLocus I)ᶜ = s := by rw [isCompact_isOpen_iff] exact ⟨fun ⟨s, e⟩ ↦ ⟨.span s, ⟨s, rfl⟩, by simpa using e⟩, fun ⟨I, ⟨s, hs⟩, e⟩ ↦ ⟨s, by simpa [hs.symm] using e⟩⟩ -lemma exists_idempotent_basicOpen_eq_of_is_clopen {s : Set (PrimeSpectrum R)} - (hs : IsClopen s) : ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := by +lemma basicOpen_injOn_isIdempotentElem : + {e : R | IsIdempotentElem e}.InjOn basicOpen := fun x hx y hy eq ↦ by + by_contra! ne + wlog ne' : x * y ≠ x generalizing x y + · apply this y hy x hx eq.symm ne.symm + rwa [mul_comm, of_not_not ne'] + have : x ∉ Ideal.span {y} := fun mem ↦ ne' <| by + obtain ⟨r, rfl⟩ := Ideal.mem_span_singleton'.mp mem + rw [mul_assoc, hy] + have ⟨p, prime, le, nmem⟩ := Ideal.exists_le_prime_nmem_of_isIdempotentElem _ x hx this + exact ne_of_mem_of_not_mem' (a := ⟨p, prime⟩) nmem + (not_not.mpr <| p.span_singleton_le_iff_mem.mp le) eq + +@[stacks 00EE] +lemma existsUnique_idempotent_basicOpen_eq_of_isClopen {s : Set (PrimeSpectrum R)} + (hs : IsClopen s) : ∃! e : R, IsIdempotentElem e ∧ s = basicOpen e := by + refine exists_unique_of_exists_of_unique ?_ ?_; swap + · rintro x y ⟨hx, rfl⟩ ⟨hy, eq⟩ + exact basicOpen_injOn_isIdempotentElem hx hy (SetLike.ext' eq) cases subsingleton_or_nontrivial R · exact ⟨0, Subsingleton.elim _ _, Subsingleton.elim _ _⟩ obtain ⟨I, hI, hI'⟩ := isCompact_isOpen_iff_ideal.mp ⟨hs.1.isCompact, hs.2⟩ @@ -712,6 +769,13 @@ 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) +lemma exists_idempotent_basicOpen_eq_of_isClopen {s : Set (PrimeSpectrum R)} + (hs : IsClopen s) : ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := + (existsUnique_idempotent_basicOpen_eq_of_isClopen hs).exists + +@[deprecated (since := "2024-11-11")] +alias exists_idempotent_basicOpen_eq_of_is_clopen := exists_idempotent_basicOpen_eq_of_isClopen + section IsIntegral open Polynomial @@ -939,7 +1003,10 @@ section Idempotent variable {R} [CommRing R] -lemma PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem +namespace PrimeSpectrum + +@[stacks 00EC] +lemma basicOpen_eq_zeroLocus_of_isIdempotentElem (e : R) (he : IsIdempotentElem e) : basicOpen e = zeroLocus {1 - e} := by ext p @@ -951,24 +1018,123 @@ lemma PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem rw [Ideal.eq_top_iff_one, ← sub_add_cancel 1 e] exact add_mem h₁ h₂ -lemma PrimeSpectrum.zeroLocus_eq_basicOpen_of_isIdempotentElem +@[stacks 00EC] +lemma zeroLocus_eq_basicOpen_of_isIdempotentElem (e : R) (he : IsIdempotentElem e) : zeroLocus {e} = basicOpen (1 - e) := by rw [basicOpen_eq_zeroLocus_of_isIdempotentElem _ he.one_sub, sub_sub_cancel] -lemma PrimeSpectrum.isClopen_iff {s : Set (PrimeSpectrum R)} : +lemma isClopen_iff {s : Set (PrimeSpectrum R)} : IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = basicOpen e := by - refine ⟨PrimeSpectrum.exists_idempotent_basicOpen_eq_of_is_clopen, ?_⟩ + refine ⟨exists_idempotent_basicOpen_eq_of_isClopen, ?_⟩ rintro ⟨e, he, rfl⟩ refine ⟨?_, (basicOpen e).2⟩ rw [PrimeSpectrum.basicOpen_eq_zeroLocus_of_isIdempotentElem e he] exact isClosed_zeroLocus _ -lemma PrimeSpectrum.isClopen_iff_zeroLocus {s : Set (PrimeSpectrum R)} : - IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = zeroLocus {e} := by - rw [isClopen_iff] - refine ⟨fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, - h.trans (basicOpen_eq_zeroLocus_of_isIdempotentElem e he)⟩, fun ⟨e, he, h⟩ ↦ - ⟨1 - e, he.one_sub, h.trans (zeroLocus_eq_basicOpen_of_isIdempotentElem e he)⟩⟩ +lemma isClopen_iff_zeroLocus {s : Set (PrimeSpectrum R)} : + IsClopen s ↔ ∃ e : R, IsIdempotentElem e ∧ s = zeroLocus {e} := + isClopen_iff.trans <| ⟨fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, + h.trans (basicOpen_eq_zeroLocus_of_isIdempotentElem e he)⟩, + fun ⟨e, he, h⟩ ↦ ⟨1 - e, he.one_sub, h.trans (zeroLocus_eq_basicOpen_of_isIdempotentElem e he)⟩⟩ + +open TopologicalSpace (Clopens Opens) + +/-- Clopen subsets in the prime spectrum of a commutative ring are in 1-1 correspondence +with idempotent elements in the ring. -/ +@[stacks 00EE] +def isIdempotentElemEquivClopens : + {e : R | IsIdempotentElem e} ≃ Clopens (PrimeSpectrum R) := + .ofBijective (fun e ↦ ⟨basicOpen e.1, isClopen_iff.mpr ⟨_, e.2, rfl⟩⟩) + ⟨fun x y eq ↦ Subtype.ext (basicOpen_injOn_isIdempotentElem x.2 y.2 <| + SetLike.ext' (congr_arg (·.1) eq)), fun s ↦ + have ⟨e, he, h⟩ := exists_idempotent_basicOpen_eq_of_isClopen s.2 + ⟨⟨e, he⟩, Clopens.ext h.symm⟩⟩ + +lemma basicOpen_isIdempotentElemEquivClopens_symm (s) : + basicOpen (isIdempotentElemEquivClopens (R := R).symm s).1 = s.toOpens := + Opens.ext <| congr_arg (·.1) (isIdempotentElemEquivClopens.apply_symm_apply s) + +lemma coe_isIdempotentElemEquivClopens_apply (e) : + (isIdempotentElemEquivClopens e : Set (PrimeSpectrum R)) = basicOpen (e.1 : R) := rfl + +lemma isIdempotentElemEquivClopens_apply_toOpens (e) : + (isIdempotentElemEquivClopens e).toOpens = basicOpen (e.1 : R) := rfl + +lemma isIdempotentElemEquivClopens_mul (e₁ e₂ : {e : R | IsIdempotentElem e}) : + isIdempotentElemEquivClopens ⟨_, e₁.2.mul e₂.2⟩ = + isIdempotentElemEquivClopens e₁ ⊓ isIdempotentElemEquivClopens e₂ := + Clopens.ext <| by simp_rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_mul]; rfl + +lemma isIdempotentElemEquivClopens_one_sub (e : {e : R | IsIdempotentElem e}) : + isIdempotentElemEquivClopens ⟨_, e.2.one_sub⟩ = (isIdempotentElemEquivClopens e)ᶜ := + SetLike.ext' <| by + simp_rw [Clopens.coe_compl, coe_isIdempotentElemEquivClopens_apply] + rw [basicOpen_eq_zeroLocus_compl, basicOpen_eq_zeroLocus_of_isIdempotentElem _ e.2] + +lemma isIdempotentElemEquivClopens_symm_inf (s₁ s₂) : + letI e := isIdempotentElemEquivClopens (R := R).symm + e (s₁ ⊓ s₂) = ⟨_, (e s₁).2.mul (e s₂).2⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| by + simp_rw [isIdempotentElemEquivClopens_mul, Equiv.apply_symm_apply] + +lemma isIdempotentElemEquivClopens_symm_compl (s : Clopens (PrimeSpectrum R)) : + isIdempotentElemEquivClopens.symm sᶜ = ⟨_, (isIdempotentElemEquivClopens.symm s).2.one_sub⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| by + rw [isIdempotentElemEquivClopens_one_sub, Equiv.apply_symm_apply] + +lemma isIdempotentElemEquivClopens_symm_top : + isIdempotentElemEquivClopens.symm ⊤ = ⟨(1 : R), .one⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| Clopens.ext <| by + rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_one]; rfl + +lemma isIdempotentElemEquivClopens_symm_bot : + isIdempotentElemEquivClopens.symm ⊥ = ⟨(0 : R), .zero⟩ := + isIdempotentElemEquivClopens.symm_apply_eq.mpr <| Clopens.ext <| by + rw [coe_isIdempotentElemEquivClopens_apply, basicOpen_zero]; rfl + +lemma isIdempotentElemEquivClopens_symm_sup (s₁ s₂ : Clopens (PrimeSpectrum R)) : + letI e := isIdempotentElemEquivClopens (R := R).symm + e (s₁ ⊔ s₂) = ⟨_, (e s₁).2.add_sub_mul (e s₂).2⟩ := Subtype.ext <| by + rw [← compl_compl (_ ⊔ _), compl_sup, isIdempotentElemEquivClopens_symm_compl] + simp_rw [isIdempotentElemEquivClopens_symm_inf, isIdempotentElemEquivClopens_symm_compl] + ring + +end PrimeSpectrum + +variable [DiscreteTopology (PrimeSpectrum R)] +open PrimeSpectrum + +variable (R) in +lemma RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : + Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by + let idem I := isIdempotentElemEquivClopens (R := R).symm ⟨{I}, isClopen_discrete _⟩ + let ideal I := Ideal.span {1 - (idem I).1} + let toSpec (I : {I : Ideal R | I.IsMaximal}) : PrimeSpectrum R := ⟨I.1, I.2.isPrime⟩ + have loc I : IsLocalization.AtPrime (R ⧸ ideal I) I.1 := by + rw [← isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton] + exacts [IsLocalization.Away.quotient_of_isIdempotentElem (idem I).2, + congr_arg (·.1) (basicOpen_isIdempotentElemEquivClopens_symm ⟨{I}, isClopen_discrete _⟩)] + let equiv I := IsLocalization.algEquiv I.1.primeCompl (Localization.AtPrime I.1) (R ⧸ ideal I) + have := (discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical.mp ‹_›).1 + have ⟨r, hr⟩ := Ideal.pi_quotient_surjective ?_ fun I ↦ equiv (toSpec I) (x I) + · refine ⟨r, funext fun I ↦ (equiv <| toSpec I).injective ?_⟩ + rw [← hr]; exact (equiv _).commutes r + refine fun I J ne ↦ Ideal.isCoprime_iff_exists.mpr ?_ + have := ((idem <| toSpec I).2.mul (idem <| toSpec J).2).eq_zero_of_isNilpotent <| by + simp_rw [← basicOpen_eq_bot_iff, basicOpen_mul, SetLike.ext'_iff, idem, + TopologicalSpace.Opens.coe_inf, basicOpen_isIdempotentElemEquivClopens_symm] + exact Set.singleton_inter_eq_empty.mpr fun h ↦ ne (Subtype.ext <| congr_arg (·.1) h) + simp_rw [ideal, Ideal.mem_span_singleton', exists_exists_eq_and] + exact ⟨1, idem (toSpec I), by simpa [mul_sub]⟩ + +/-- If the prime spectrum of a commutative ring R has discrete Zariski topology, then R is +canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ +@[stacks 00JA +"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] +def RingHom.toLocalizationIsMaximalEquiv : R ≃+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, + RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ end Idempotent diff --git a/Mathlib/RingTheory/Artinian.lean b/Mathlib/RingTheory/Artinian.lean index f61383992490d..2911d10aa8bb9 100644 --- a/Mathlib/RingTheory/Artinian.lean +++ b/Mathlib/RingTheory/Artinian.lean @@ -10,6 +10,7 @@ import Mathlib.Order.Filter.EventuallyConst import Mathlib.RingTheory.Nakayama import Mathlib.RingTheory.SimpleModule import Mathlib.Tactic.RSuffices +import Mathlib.Tactic.StacksAttribute /-! # Artinian rings and modules @@ -347,6 +348,7 @@ end CommRing Strictly speaking, this should be called `IsLeftArtinianRing` but we omit the `Left` for convenience in the commutative case. For a right Artinian ring, use `IsArtinian Rᵐᵒᵖ R`. -/ +@[stacks 00J5] abbrev IsArtinianRing (R) [Ring R] := IsArtinian R R @@ -418,6 +420,7 @@ open IsArtinian variable {R : Type*} [CommRing R] [IsArtinianRing R] +@[stacks 00J8] theorem isNilpotent_jacobson_bot : IsNilpotent (Ideal.jacobson (⊥ : Ideal R)) := by let Jac := Ideal.jacobson (⊥ : Ideal R) let f : ℕ →o (Ideal R)ᵒᵈ := ⟨fun n => Jac ^ n, fun _ _ h => Ideal.pow_le_pow_right h⟩ @@ -504,9 +507,7 @@ lemma primeSpectrum_finite : {I : Ideal R | I.IsPrime}.Finite := by rwa [← Subtype.ext <| (@isMaximal_of_isPrime _ _ _ _ q.2).eq_of_le p.2.1 hq2] variable (R) -/-- -[Stacks Lemma 00J7](https://stacks.math.columbia.edu/tag/00J7) --/ +@[stacks 00J7] lemma maximal_ideals_finite : {I : Ideal R | I.IsMaximal}.Finite := by simp_rw [← isPrime_iff_isMaximal] apply primeSpectrum_finite R diff --git a/Mathlib/RingTheory/Ideal/Maximal.lean b/Mathlib/RingTheory/Ideal/Maximal.lean index 444898cdd4f69..c3cb20cf054d0 100644 --- a/Mathlib/RingTheory/Ideal/Maximal.lean +++ b/Mathlib/RingTheory/Ideal/Maximal.lean @@ -200,6 +200,26 @@ lemma isPrime_of_maximally_disjoint (I : Ideal α) rw [← hr₁, ← hr₂] ring +theorem exists_le_prime_disjoint (S : Submonoid α) (disjoint : Disjoint (I : Set α) S) : + ∃ p : Ideal α, p.IsPrime ∧ I ≤ p ∧ Disjoint (p : Set α) S := by + have ⟨p, hIp, hp⟩ := zorn_le_nonempty₀ {p : Ideal α | Disjoint (p : Set α) S} + (fun c hc hc' x hx ↦ ?_) I disjoint + · exact ⟨p, isPrime_of_maximally_disjoint _ _ hp.1 (fun _ ↦ hp.not_prop_of_gt), hIp, hp.1⟩ + cases isEmpty_or_nonempty c + · exact ⟨I, disjoint, fun J hJ ↦ isEmptyElim (⟨J, hJ⟩ : c)⟩ + refine ⟨sSup c, Set.disjoint_left.mpr fun x hx ↦ ?_, fun _ ↦ le_sSup⟩ + have ⟨p, hp⟩ := (Submodule.mem_iSup_of_directed _ hc'.directed).mp (sSup_eq_iSup' c ▸ hx) + exact Set.disjoint_left.mp (hc p.2) hp + +theorem exists_le_prime_nmem_of_isIdempotentElem (a : α) (ha : IsIdempotentElem a) (haI : a ∉ I) : + ∃ p : Ideal α, p.IsPrime ∧ I ≤ p ∧ a ∉ p := + have : Disjoint (I : Set α) (Submonoid.powers a) := Set.disjoint_right.mpr <| by + rw [ha.coe_powers] + rintro _ (rfl|rfl) + exacts [I.ne_top_iff_one.mp (ne_of_mem_of_not_mem' Submodule.mem_top haI).symm, haI] + have ⟨p, h1, h2, h3⟩ := exists_le_prime_disjoint _ _ this + ⟨p, h1, h2, Set.disjoint_right.mp h3 (Submonoid.mem_powers a)⟩ + end Ideal end CommSemiring diff --git a/Mathlib/RingTheory/Ideal/Span.lean b/Mathlib/RingTheory/Ideal/Span.lean index e2f263723ce31..1b10889acf12f 100644 --- a/Mathlib/RingTheory/Ideal/Span.lean +++ b/Mathlib/RingTheory/Ideal/Span.lean @@ -237,3 +237,22 @@ theorem span_singleton_neg (x : α) : (span {-x} : Ideal α) = span {x} := by end Ideal end Ring + +namespace IsIdempotentElem + +variable {R} [CommRing R] {e : R} (he : IsIdempotentElem e) +include he + +theorem ker_toSpanSingleton_eq_span : + LinearMap.ker (LinearMap.toSpanSingleton R R e) = Ideal.span {1 - e} := SetLike.ext fun x ↦ by + rw [Ideal.mem_span_singleton'] + refine ⟨fun h ↦ ⟨x, by rw [mul_sub, show x * e = 0 from h, mul_one, sub_zero]⟩, fun h ↦ ?_⟩ + obtain ⟨x, rfl⟩ := h + show x * (1 - e) * e = 0 + rw [mul_assoc, sub_mul, one_mul, he, sub_self, mul_zero] + +theorem ker_toSpanSingleton_one_sub_eq_span : + LinearMap.ker (LinearMap.toSpanSingleton R R (1 - e)) = Ideal.span {e} := by + rw [ker_toSpanSingleton_eq_span he.one_sub, sub_sub_cancel] + +end IsIdempotentElem diff --git a/Mathlib/RingTheory/Idempotents.lean b/Mathlib/RingTheory/Idempotents.lean index ffc85853ec287..65b8c4836a131 100644 --- a/Mathlib/RingTheory/Idempotents.lean +++ b/Mathlib/RingTheory/Idempotents.lean @@ -7,6 +7,7 @@ import Mathlib.Algebra.GeomSum import Mathlib.Algebra.Polynomial.AlgebraMap import Mathlib.RingTheory.Ideal.Quotient.Operations import Mathlib.RingTheory.Nilpotent.Defs +import Mathlib.Tactic.StacksAttribute /-! @@ -346,6 +347,7 @@ theorem eq_of_isNilpotent_sub_of_isIdempotentElem {e₁ e₂ : R} e₁ = e₂ := eq_of_isNilpotent_sub_of_isIdempotentElem_of_commute he₁ he₂ H (.all _ _) +@[stacks 00J9] theorem existsUnique_isIdempotentElem_eq_of_ker_isNilpotent (h : ∀ x ∈ RingHom.ker f, IsNilpotent x) (e : S) (he : e ∈ f.range) (he' : IsIdempotentElem e) : ∃! e' : R, IsIdempotentElem e' ∧ f e' = e := by diff --git a/Mathlib/RingTheory/Localization/AtPrime.lean b/Mathlib/RingTheory/Localization/AtPrime.lean index 76d9a9d2cb2e8..1337c943de11e 100644 --- a/Mathlib/RingTheory/Localization/AtPrime.lean +++ b/Mathlib/RingTheory/Localization/AtPrime.lean @@ -250,3 +250,23 @@ theorem localRingHom_comp {S : Type*} [CommSemiring S] (J : Ideal S) [hJ : J.IsP simp only [Function.comp_apply, RingHom.coe_comp, localRingHom_to_map] end Localization + +namespace RingHom + +variable (R) + +/-- The canonical ring homomorphism from a commutative semiring to the product of its +localizations at all maximal ideals. It is always injective. -/ +def toLocalizationIsMaximal : R →+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + Pi.ringHom fun _ ↦ algebraMap R _ + +theorem toLocalizationIsMaximal_injective : + Function.Injective (RingHom.toLocalizationIsMaximal R) := fun r r' eq ↦ by + rw [← one_mul r, ← one_mul r'] + by_contra ne + have ⟨I, mI, hI⟩ := (Module.eqIdeal R r r').exists_le_maximal ((Ideal.ne_top_iff_one _).mpr ne) + have ⟨s, hs⟩ := (IsLocalization.eq_iff_exists I.primeCompl _).mp (congr_fun eq ⟨I, mI⟩) + exact s.2 (hI hs) + +end RingHom diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index 01a01a331e08b..ae069a4b51072 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -297,53 +297,42 @@ end AtUnits section Prod -lemma away_of_isIdempotentElem {R S} [CommRing R] [CommRing S] [Algebra R S] +lemma away_of_isIdempotentElem_of_mul {R S} [CommSemiring R] [CommSemiring S] [Algebra R S] {e : R} (he : IsIdempotentElem e) - (H : RingHom.ker (algebraMap R S) = Ideal.span {1 - e}) + (H : ∀ x y, algebraMap R S x = algebraMap R S y ↔ e * x = e * y) (H' : Function.Surjective (algebraMap R S)) : IsLocalization.Away e S where map_units' r := by - have : algebraMap R S e = 1 := by - rw [← (algebraMap R S).map_one, eq_comm, ← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, - H, Ideal.mem_span_singleton] obtain ⟨r, n, rfl⟩ := r - simp [this] + simp [show algebraMap R S e = 1 by rw [← (algebraMap R S).map_one, H, mul_one, he]] surj' z := by obtain ⟨z, rfl⟩ := H' z exact ⟨⟨z, 1⟩, by simp⟩ - exists_of_eq {x y} h := by - rw [← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, H, Ideal.mem_span_singleton] at h - obtain ⟨k, hk⟩ := h - refine ⟨⟨e, Submonoid.mem_powers e⟩, ?_⟩ - rw [← sub_eq_zero, ← mul_sub, hk, ← mul_assoc, mul_sub, mul_one, he.eq, sub_self, zero_mul] + exists_of_eq {x y} h := ⟨⟨e, Submonoid.mem_powers e⟩, (H x y).mp h⟩ -instance away_fst {R S} [CommRing R] [CommRing S] : +lemma away_of_isIdempotentElem {R S} [CommRing R] [CommRing S] [Algebra R S] + {e : R} (he : IsIdempotentElem e) + (H : RingHom.ker (algebraMap R S) = Ideal.span {1 - e}) + (H' : Function.Surjective (algebraMap R S)) : + IsLocalization.Away e S := + away_of_isIdempotentElem_of_mul he (fun x y ↦ by + rw [← sub_eq_zero, ← map_sub, ← RingHom.mem_ker, H, ← he.ker_toSpanSingleton_eq_span, + LinearMap.mem_ker, LinearMap.toSpanSingleton_apply, sub_smul, sub_eq_zero] + simp_rw [mul_comm e, smul_eq_mul]) H' + +instance away_fst {R S} [CommSemiring R] [CommSemiring S] : letI := (RingHom.fst R S).toAlgebra - IsLocalization.Away (R := R × S) (1, 0) R := by + IsLocalization.Away (R := R × S) (1, 0) R := letI := (RingHom.fst R S).toAlgebra - apply away_of_isIdempotentElem - · ext <;> simp - · ext x - simp only [RingHom.algebraMap_toAlgebra, RingHom.mem_ker, RingHom.coe_fst, - Ideal.mem_span_singleton, Prod.one_eq_mk, Prod.mk_sub_mk, sub_self, sub_zero] - constructor - · intro e; use x; ext <;> simp [e] - · rintro ⟨⟨i, j⟩, rfl⟩; simp - · exact Prod.fst_surjective - -instance away_snd {R S} [CommRing R] [CommRing S] : + away_of_isIdempotentElem_of_mul (by ext <;> simp) + (fun ⟨xR, xS⟩ ⟨yR, yS⟩ ↦ show xR = yR ↔ _ by simp) Prod.fst_surjective + +instance away_snd {R S} [CommSemiring R] [CommSemiring S] : letI := (RingHom.snd R S).toAlgebra - IsLocalization.Away (R := R × S) (0, 1) S := by + IsLocalization.Away (R := R × S) (0, 1) S := letI := (RingHom.snd R S).toAlgebra - apply away_of_isIdempotentElem - · ext <;> simp - · ext x - simp only [RingHom.algebraMap_toAlgebra, RingHom.mem_ker, RingHom.coe_snd, - Ideal.mem_span_singleton, Prod.one_eq_mk, Prod.mk_sub_mk, sub_self, sub_zero] - constructor - · intro e; use x; ext <;> simp [e] - · rintro ⟨⟨i, j⟩, rfl⟩; simp - · exact Prod.snd_surjective + away_of_isIdempotentElem_of_mul (by ext <;> simp) + (fun ⟨xR, xS⟩ ⟨yR, yS⟩ ↦ show xS = yS ↔ _ by simp) Prod.snd_surjective end Prod diff --git a/Mathlib/RingTheory/Localization/Away/Lemmas.lean b/Mathlib/RingTheory/Localization/Away/Lemmas.lean index 640763c874cbe..dc5d5600d2f7f 100644 --- a/Mathlib/RingTheory/Localization/Away/Lemmas.lean +++ b/Mathlib/RingTheory/Localization/Away/Lemmas.lean @@ -15,9 +15,7 @@ This file contains lemmas on localization away from an element requiring more im variable {R : Type*} [CommRing R] -namespace IsLocalization - -namespace Away +namespace IsLocalization.Away /-- Given a set `s` in a ring `R` and for every `t : s` a set `p t` of fractions in a localization of `R` at `t`, this is the function sending a pair `(t, y)`, with @@ -61,6 +59,8 @@ lemma span_range_mulNumerator_eq_top {s : Set R} use n + m simpa [pow_add, hc] using Ideal.mul_mem_left _ _ hy -end Away +lemma quotient_of_isIdempotentElem {e : R} (he : IsIdempotentElem e) : + IsLocalization.Away e (R ⧸ Ideal.span {1 - e}) := + away_of_isIdempotentElem he Ideal.mk_ker Quotient.mk_surjective -end IsLocalization +end IsLocalization.Away diff --git a/Mathlib/RingTheory/Localization/Basic.lean b/Mathlib/RingTheory/Localization/Basic.lean index f2d6adffdcf45..9159d55e9cdbd 100644 --- a/Mathlib/RingTheory/Localization/Basic.lean +++ b/Mathlib/RingTheory/Localization/Basic.lean @@ -73,7 +73,7 @@ open Function section CommSemiring -variable {R : Type*} [CommSemiring R] {M : Submonoid R} {S : Type*} [CommSemiring S] +variable {R : Type*} [CommSemiring R] {M N : Submonoid R} {S : Type*} [CommSemiring S] variable [Algebra R S] {P : Type*} [CommSemiring P] namespace IsLocalization @@ -237,7 +237,7 @@ end IsLocalization section -variable (M) +variable (M N) theorem isLocalization_of_algEquiv [Algebra R P] [IsLocalization M S] (h : S ≃ₐ[R] P) : IsLocalization M P := by @@ -265,6 +265,14 @@ theorem isLocalization_iff_of_ringEquiv (h : S ≃+* P) : letI := (h.toRingHom.comp <| algebraMap R S).toAlgebra isLocalization_iff_of_algEquiv M { h with commutes' := fun _ => rfl } +variable (S) in +/-- If an algebra is simultaneously localizations for two submonoids, then an arbitrary algebra +is a localization of one submonoid iff it is a localization of the other. -/ +theorem isLocalization_iff_of_isLocalization [IsLocalization M S] [IsLocalization N S] + [Algebra R P] : IsLocalization M P ↔ IsLocalization N P := + ⟨fun _ ↦ isLocalization_of_algEquiv N (algEquiv M S P), + fun _ ↦ isLocalization_of_algEquiv M (algEquiv N S P)⟩ + end variable (M) diff --git a/Mathlib/RingTheory/Localization/Defs.lean b/Mathlib/RingTheory/Localization/Defs.lean index 486cd6a8eb03e..60509638366d3 100644 --- a/Mathlib/RingTheory/Localization/Defs.lean +++ b/Mathlib/RingTheory/Localization/Defs.lean @@ -140,6 +140,11 @@ theorem of_le (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ r ∈ N, IsUnit (al rintro ⟨c, hc⟩ exact ⟨⟨c, h₁ c.2⟩, hc⟩ +theorem of_le_of_exists_dvd (N : Submonoid R) (h₁ : M ≤ N) (h₂ : ∀ n ∈ N, ∃ m ∈ M, n ∣ m) : + IsLocalization N S := + of_le M N h₁ fun n hn ↦ have ⟨m, hm, dvd⟩ := h₂ n hn + isUnit_of_dvd_unit (map_dvd _ dvd) (map_units S ⟨m, hm⟩) + variable (S) /-- `IsLocalization.toLocalizationWithZeroMap M S` shows `S` is the monoid localization of diff --git a/Mathlib/RingTheory/Nilpotent/Basic.lean b/Mathlib/RingTheory/Nilpotent/Basic.lean index 9eddb6417010a..602de4bcb8dbd 100644 --- a/Mathlib/RingTheory/Nilpotent/Basic.lean +++ b/Mathlib/RingTheory/Nilpotent/Basic.lean @@ -91,6 +91,14 @@ lemma IsNilpotent.not_isUnit [Ring R] [Nontrivial R] {x : R} (hx : IsNilpotent x ¬ IsUnit x := mt IsUnit.not_isNilpotent (by simpa only [not_not] using hx) +lemma IsIdempotentElem.eq_zero_of_isNilpotent [MonoidWithZero R] {e : R} + (idem : IsIdempotentElem e) (nilp : IsNilpotent e) : e = 0 := by + obtain ⟨rfl | n, hn⟩ := nilp + · rw [pow_zero] at hn; rw [← one_mul e, hn, zero_mul] + · rw [← hn, idem.pow_succ_eq] + +alias IsNilpotent.eq_zero_of_isIdempotentElem := IsIdempotentElem.eq_zero_of_isNilpotent + instance [Zero R] [Pow R ℕ] [Zero S] [Pow S ℕ] [IsReduced R] [IsReduced S] : IsReduced (R × S) where eq_zero _ := fun ⟨n, hn⟩ ↦ have hn := Prod.ext_iff.1 hn Prod.ext (IsReduced.eq_zero _ ⟨n, hn.1⟩) (IsReduced.eq_zero _ ⟨n, hn.2⟩) diff --git a/Mathlib/RingTheory/PrimeSpectrum.lean b/Mathlib/RingTheory/PrimeSpectrum.lean index dd7f6aa053423..5ac44c10a7f5f 100644 --- a/Mathlib/RingTheory/PrimeSpectrum.lean +++ b/Mathlib/RingTheory/PrimeSpectrum.lean @@ -83,6 +83,14 @@ instance [Subsingleton R] : IsEmpty (PrimeSpectrum R) := variable (R S) +/-- The prime spectrum is in bijection with the set of prime ideals. -/ +@[simps] +def equivSubtype : PrimeSpectrum R ≃ {I : Ideal R // I.IsPrime} where + toFun I := ⟨I.asIdeal, I.2⟩ + invFun I := ⟨I, I.2⟩ + left_inv _ := rfl + right_inv _ := rfl + /-- The map from the direct sum of prime spectra to the prime spectrum of a direct product. -/ @[simp] def primeSpectrumProdOfSum : PrimeSpectrum R ⊕ PrimeSpectrum S → PrimeSpectrum (R × S) diff --git a/Mathlib/Topology/Order.lean b/Mathlib/Topology/Order.lean index 789c64a0b0511..168716335f5dc 100644 --- a/Mathlib/Topology/Order.lean +++ b/Mathlib/Topology/Order.lean @@ -305,6 +305,11 @@ theorem singletons_open_iff_discrete {X : Type*} [TopologicalSpace X] : (∀ a : X, IsOpen ({a} : Set X)) ↔ DiscreteTopology X := ⟨fun h => ⟨eq_bot_of_singletons_open h⟩, fun a _ => @isOpen_discrete _ _ a _⟩ +theorem DiscreteTopology.of_finite_of_isClosed_singleton [TopologicalSpace α] [Finite α] + (h : ∀ a : α, IsClosed {a}) : DiscreteTopology α := + discreteTopology_iff_forall_isClosed.mpr fun s ↦ + s.iUnion_of_singleton_coe ▸ isClosed_iUnion_of_finite fun _ ↦ h _ + theorem discreteTopology_iff_singleton_mem_nhds [TopologicalSpace α] : DiscreteTopology α ↔ ∀ x : α, {x} ∈ 𝓝 x := by simp only [← singletons_open_iff_discrete, isOpen_iff_mem_nhds, mem_singleton_iff, forall_eq] From 5c08dca69b1335bc2a146801d76c98a56bc24724 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:02:00 +0000 Subject: [PATCH 153/250] feat: `Finset.sup` in a group (#19460) Co-authored-by: Andrew Yang --- Mathlib.lean | 1 + Mathlib/Algebra/Order/Group/Finset.lean | 81 ++++++++++++++++++++ Mathlib/Algebra/Polynomial/Basic.lean | 3 +- Mathlib/Data/Finset/Fold.lean | 7 -- Mathlib/Data/Finset/Lattice/Fold.lean | 20 ++--- Mathlib/Order/Filter/AtTopBot/Monoid.lean | 1 + Mathlib/RingTheory/HahnSeries/Basic.lean | 1 + Mathlib/Topology/Algebra/ConstMulAction.lean | 1 + 8 files changed, 92 insertions(+), 23 deletions(-) create mode 100644 Mathlib/Algebra/Order/Group/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index c6a763cc2c95e..344f1930eae20 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -655,6 +655,7 @@ import Mathlib.Algebra.Order.Group.CompleteLattice import Mathlib.Algebra.Order.Group.Cone import Mathlib.Algebra.Order.Group.Defs import Mathlib.Algebra.Order.Group.DenselyOrdered +import Mathlib.Algebra.Order.Group.Finset import Mathlib.Algebra.Order.Group.Indicator import Mathlib.Algebra.Order.Group.InjSurj import Mathlib.Algebra.Order.Group.Instances diff --git a/Mathlib/Algebra/Order/Group/Finset.lean b/Mathlib/Algebra/Order/Group/Finset.lean new file mode 100644 index 0000000000000..e08078aa0b808 --- /dev/null +++ b/Mathlib/Algebra/Order/Group/Finset.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2024 Yaël Dillies, Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yaël Dillies, Andrew Yang +-/ +import Mathlib.Algebra.Order.Group.OrderIso +import Mathlib.Algebra.Order.Monoid.Canonical.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop +import Mathlib.Data.Finset.Lattice.Fold + +/-! +# `Finset.sup` in a group +-/ + +assert_not_exists MonoidWithZero + +namespace Finset +variable {ι κ M G : Type*} + +lemma fold_max_add [LinearOrder M] [Add M] [AddRightMono M] (s : Finset ι) (a : WithBot M) + (f : ι → M) : s.fold max ⊥ (fun i ↦ ↑(f i) + a) = s.fold max ⊥ ((↑) ∘ f) + a := by + classical + induction' s using Finset.induction_on with a s _ ih <;> simp [*, max_add_add_right] + +@[to_additive nsmul_inf'] +lemma inf'_pow [LinearOrder M] [Monoid M] [MulLeftMono M] [MulRightMono M] (s : Finset ι) + (f : ι → M) (n : ℕ) (hs) : s.inf' hs f ^ n = s.inf' hs fun a ↦ f a ^ n := + map_finset_inf' (OrderHom.mk _ <| pow_left_mono n) hs _ + +@[to_additive nsmul_sup'] +lemma sup'_pow [LinearOrder M] [Monoid M] [MulLeftMono M] [MulRightMono M] (s : Finset ι) + (f : ι → M) (n : ℕ) (hs) : s.sup' hs f ^ n = s.sup' hs fun a ↦ f a ^ n := + map_finset_sup' (OrderHom.mk _ <| pow_left_mono n) hs _ + +section Group +variable [Group G] [LinearOrder G] + +@[to_additive "Also see `Finset.sup'_add'` that works for canonically ordered monoids."] +lemma sup'_mul [MulRightMono G] (s : Finset ι) (f : ι → G) (a : G) (hs) : + s.sup' hs f * a = s.sup' hs fun i ↦ f i * a := map_finset_sup' (OrderIso.mulRight a) hs f + +set_option linter.docPrime false in +@[to_additive "Also see `Finset.add_sup''` that works for canonically ordered monoids."] +lemma mul_sup' [MulLeftMono G] (s : Finset ι) (f : ι → G) (a : G) (hs) : + a * s.sup' hs f = s.sup' hs fun i ↦ a * f i := map_finset_sup' (OrderIso.mulLeft a) hs f + +end Group + +section CanonicallyLinearOrderedAddCommMonoid +variable [CanonicallyLinearOrderedAddCommMonoid M] [Sub M] [AddLeftReflectLE M] [OrderedSub M] + {s : Finset ι} {t : Finset κ} + +/-- Also see `Finset.sup'_add` that works for ordered groups. -/ +lemma sup'_add' (s : Finset ι) (f : ι → M) (a : M) (hs : s.Nonempty) : + s.sup' hs f + a = s.sup' hs fun i ↦ f i + a := by + apply le_antisymm + · apply add_le_of_le_tsub_right_of_le + · exact Finset.le_sup'_of_le _ hs.choose_spec le_add_self + · exact Finset.sup'_le _ _ fun i hi ↦ le_tsub_of_add_le_right (Finset.le_sup' (f · + a) hi) + · exact Finset.sup'_le _ _ fun i hi ↦ add_le_add_right (Finset.le_sup' _ hi) _ + +/-- Also see `Finset.add_sup'` that works for ordered groups. -/ +lemma add_sup'' (hs : s.Nonempty) (f : ι → M) (a : M) : + a + s.sup' hs f = s.sup' hs fun i ↦ a + f i := by simp_rw [add_comm a, Finset.sup'_add'] + +protected lemma sup_add (hs : s.Nonempty) (f : ι → M) (a : M) : + s.sup f + a = s.sup fun i ↦ f i + a := by + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, sup'_add'] + +protected lemma add_sup (hs : s.Nonempty) (f : ι → M) (a : M) : + a + s.sup f = s.sup fun i ↦ a + f i := by + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, add_sup''] + +lemma sup_add_sup (hs : s.Nonempty) (ht : t.Nonempty) (f : ι → M) (g : κ → M) : + s.sup f + t.sup g = (s ×ˢ t).sup fun ij ↦ f ij.1 + g ij.2 := by + simp only [Finset.sup_add hs, Finset.add_sup ht, Finset.sup_product_left] + +end CanonicallyLinearOrderedAddCommMonoid +end Finset diff --git a/Mathlib/Algebra/Polynomial/Basic.lean b/Mathlib/Algebra/Polynomial/Basic.lean index 56c5d26cab71c..72765a3a6e65a 100644 --- a/Mathlib/Algebra/Polynomial/Basic.lean +++ b/Mathlib/Algebra/Polynomial/Basic.lean @@ -4,8 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Johannes Hölzl, Kim Morrison, Jens Wagemaker -/ import Mathlib.Algebra.GroupWithZero.Divisibility -import Mathlib.Data.Finset.Sort import Mathlib.Algebra.MonoidAlgebra.Defs +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop +import Mathlib.Data.Finset.Sort import Mathlib.Order.OmegaCompletePartialOrder /-! diff --git a/Mathlib/Data/Finset/Fold.lean b/Mathlib/Data/Finset/Fold.lean index 931d4b86e5a6b..fb9ed13966a4d 100644 --- a/Mathlib/Data/Finset/Fold.lean +++ b/Mathlib/Data/Finset/Fold.lean @@ -3,8 +3,6 @@ Copyright (c) 2017 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import Mathlib.Algebra.Order.Monoid.Unbundled.MinMax -import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Data.Finset.Image import Mathlib.Data.Multiset.Fold @@ -231,11 +229,6 @@ theorem fold_max_lt : s.fold max b f < c ↔ b < c ∧ ∀ x ∈ s, f x < c := b theorem lt_fold_max : c < s.fold max b f ↔ c < b ∨ ∃ x ∈ s, c < f x := fold_op_rel_iff_or lt_max_iff -theorem fold_max_add [Add β] [AddRightMono β] (n : WithBot β) - (s : Finset α) : (s.fold max ⊥ fun x : α => ↑(f x) + n) = s.fold max ⊥ ((↑) ∘ f) + n := by - classical - induction' s using Finset.induction_on with a s _ ih <;> simp [*, max_add_add_right] - end Order end Fold diff --git a/Mathlib/Data/Finset/Lattice/Fold.lean b/Mathlib/Data/Finset/Lattice/Fold.lean index 4c0a9c4e8b587..404b2ac9105b3 100644 --- a/Mathlib/Data/Finset/Lattice/Fold.lean +++ b/Mathlib/Data/Finset/Lattice/Fold.lean @@ -3,7 +3,6 @@ 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.Monoid.Unbundled.Pow import Mathlib.Data.Finset.Fold import Mathlib.Data.Finset.Pi import Mathlib.Data.Finset.Prod @@ -819,13 +818,6 @@ theorem _root_.map_finset_sup' [SemilatticeSup β] [FunLike F α β] [SupHomClas f (s.sup' hs g) = s.sup' hs (f ∘ g) := by refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] -lemma nsmul_sup' {α β : Type*} [AddMonoid β] [LinearOrder β] - [AddLeftMono β] [AddRightMono β] - {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : - s.sup' hs (fun a => n • f a) = n • s.sup' hs f := - let ns : SupHom β β := { toFun := (n • ·), map_sup' := fun _ _ => (nsmul_right_mono n).map_max } - (map_finset_sup' ns hs _).symm - /-- To rewrite from right to left, use `Finset.sup'_comp_eq_image`. -/ @[simp] theorem sup'_image [DecidableEq β] {s : Finset γ} {f : γ → β} (hs : (s.image f).Nonempty) @@ -857,6 +849,11 @@ lemma sup'_comp_eq_map {s : Finset γ} {f : γ ↪ β} (g : β → α) (hs : s.N theorem sup'_mono {s₁ s₂ : Finset β} (h : s₁ ⊆ s₂) (h₁ : s₁.Nonempty) : s₁.sup' h₁ f ≤ s₂.sup' (h₁.mono h) f := Finset.sup'_le h₁ _ (fun _ hb => le_sup' _ (h hb)) + +@[gcongr] +lemma sup'_mono_fun {hs : s.Nonempty} {f g : β → α} (h : ∀ b ∈ s, f b ≤ g b) : + s.sup' hs f ≤ s.sup' hs g := sup'_le _ _ fun b hb ↦ (h b hb).trans (le_sup' _ hb) + end Sup' section Inf' @@ -974,13 +971,6 @@ theorem _root_.map_finset_inf' [SemilatticeInf β] [FunLike F α β] [InfHomClas f (s.inf' hs g) = s.inf' hs (f ∘ g) := by refine hs.cons_induction ?_ ?_ <;> intros <;> simp [*] -lemma nsmul_inf' {α β : Type*} [AddMonoid β] [LinearOrder β] - [AddLeftMono β] [AddRightMono β] - {s : Finset α} (hs : s.Nonempty) (f : α → β) (n : ℕ) : - s.inf' hs (fun a => n • f a) = n • s.inf' hs f := - let ns : InfHom β β := { toFun := (n • ·), map_inf' := fun _ _ => (nsmul_right_mono n).map_min } - (map_finset_inf' ns hs _).symm - /-- To rewrite from right to left, use `Finset.inf'_comp_eq_image`. -/ @[simp] theorem inf'_image [DecidableEq β] {s : Finset γ} {f : γ → β} (hs : (s.image f).Nonempty) diff --git a/Mathlib/Order/Filter/AtTopBot/Monoid.lean b/Mathlib/Order/Filter/AtTopBot/Monoid.lean index 0cfa275d9a426..02cdf596f5854 100644 --- a/Mathlib/Order/Filter/AtTopBot/Monoid.lean +++ b/Mathlib/Order/Filter/AtTopBot/Monoid.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ import Mathlib.Algebra.Order.Monoid.OrderDual +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Order.Filter.AtTopBot /-! diff --git a/Mathlib/RingTheory/HahnSeries/Basic.lean b/Mathlib/RingTheory/HahnSeries/Basic.lean index aa4ddc454c2dc..b074a86e1aa3b 100644 --- a/Mathlib/RingTheory/HahnSeries/Basic.lean +++ b/Mathlib/RingTheory/HahnSeries/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Aaron Anderson -/ import Mathlib.Algebra.Group.Support +import Mathlib.Algebra.Order.Monoid.Unbundled.WithTop import Mathlib.Order.WellFoundedSet /-! diff --git a/Mathlib/Topology/Algebra/ConstMulAction.lean b/Mathlib/Topology/Algebra/ConstMulAction.lean index 44b6433923f59..10007257a47c5 100644 --- a/Mathlib/Topology/Algebra/ConstMulAction.lean +++ b/Mathlib/Topology/Algebra/ConstMulAction.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, Heather Macbeth -/ import Mathlib.Algebra.Module.ULift +import Mathlib.Algebra.Order.Group.Synonym import Mathlib.Data.Set.Pointwise.SMul import Mathlib.GroupTheory.GroupAction.Defs import Mathlib.Topology.Algebra.Constructions From 25781872f1cce8492209ebd11659f3db18e9c18c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:02:02 +0000 Subject: [PATCH 154/250] feat: `Finset.sup` in a group with zero (#19461) Co-authored-by: Andrew Yang --- Mathlib.lean | 2 +- .../Algebra/Order/GroupWithZero/Finset.lean | 64 +++++++++++++++++++ Mathlib/Algebra/Order/Ring/Finset.lean | 34 ---------- Mathlib/Analysis/Normed/Field/Lemmas.lean | 4 +- 4 files changed, 67 insertions(+), 37 deletions(-) create mode 100644 Mathlib/Algebra/Order/GroupWithZero/Finset.lean delete mode 100644 Mathlib/Algebra/Order/Ring/Finset.lean diff --git a/Mathlib.lean b/Mathlib.lean index 344f1930eae20..7206497b94d35 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -679,6 +679,7 @@ import Mathlib.Algebra.Order.Group.Unbundled.Int import Mathlib.Algebra.Order.Group.Units import Mathlib.Algebra.Order.GroupWithZero.Action.Synonym import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.GroupWithZero.Finset import Mathlib.Algebra.Order.GroupWithZero.Submonoid import Mathlib.Algebra.Order.GroupWithZero.Synonym import Mathlib.Algebra.Order.GroupWithZero.Unbundled @@ -739,7 +740,6 @@ import Mathlib.Algebra.Order.Ring.Canonical import Mathlib.Algebra.Order.Ring.Cast import Mathlib.Algebra.Order.Ring.Cone import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Algebra.Order.Ring.Finset import Mathlib.Algebra.Order.Ring.InjSurj import Mathlib.Algebra.Order.Ring.Int import Mathlib.Algebra.Order.Ring.Nat diff --git a/Mathlib/Algebra/Order/GroupWithZero/Finset.lean b/Mathlib/Algebra/Order/GroupWithZero/Finset.lean new file mode 100644 index 0000000000000..0047881217ebf --- /dev/null +++ b/Mathlib/Algebra/Order/GroupWithZero/Finset.lean @@ -0,0 +1,64 @@ +/- +Copyright (c) 2022 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser, Yaël Dillies, Andrew Yang +-/ +import Mathlib.Algebra.Order.GroupWithZero.Canonical +import Mathlib.Algebra.Order.GroupWithZero.Unbundled.Lemmas +import Mathlib.Data.Finset.Lattice.Fold + +/-! +# `Finset.sup` in a group with zero +-/ + +namespace Finset +variable {ι M₀ G₀ : Type*} + +section MonoidWithZero +variable [MonoidWithZero M₀] {s : Finset ι} {a b : ι → M₀} + +lemma sup_mul_le_mul_sup_of_nonneg [SemilatticeSup M₀] [OrderBot M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup (a * b) ≤ s.sup a * s.sup b := + Finset.sup_le fun _i hi ↦ + mul_le_mul (le_sup hi) (le_sup hi) (hb _ hi) ((ha _ hi).trans <| le_sup hi) + +lemma mul_inf_le_inf_mul_of_nonneg [SemilatticeInf M₀] [OrderTop M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf a * s.inf b ≤ s.inf (a * b) := + Finset.le_inf fun i hi ↦ mul_le_mul (inf_le hi) (inf_le hi) (Finset.le_inf hb) (ha i hi) + +lemma sup'_mul_le_mul_sup'_of_nonneg [SemilatticeSup M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) (hs) : + s.sup' hs (a * b) ≤ s.sup' hs a * s.sup' hs b := + sup'_le _ _ fun _i hi ↦ + mul_le_mul (le_sup' _ hi) (le_sup' _ hi) (hb _ hi) ((ha _ hi).trans <| le_sup' _ hi) + +lemma inf'_mul_le_mul_inf'_of_nonneg [SemilatticeInf M₀] [PosMulMono M₀] [MulPosMono M₀] + (ha : ∀ i ∈ s, 0 ≤ a i) (hb : ∀ i ∈ s, 0 ≤ b i) (hs) : + s.inf' hs a * s.inf' hs b ≤ s.inf' hs (a * b) := + le_inf' _ _ fun _i hi ↦ mul_le_mul (inf'_le _ hi) (inf'_le _ hi) (le_inf' _ _ hb) (ha _ hi) + +end MonoidWithZero + +section GroupWithZero +variable [GroupWithZero G₀] [SemilatticeSup G₀] {s : Finset ι} {a : G₀} + +lemma sup'_mul₀ [MulPosMono G₀] [MulPosReflectLE G₀] (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : + s.sup' hs f * a = s.sup' hs fun i ↦ f i * a := map_finset_sup' (OrderIso.mulRight₀ _ ha) hs f + +set_option linter.docPrime false in +lemma mul₀_sup' [PosMulMono G₀] [PosMulReflectLE G₀] (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : + a * s.sup' hs f = s.sup' hs fun i ↦ a * f i := map_finset_sup' (OrderIso.mulLeft₀ _ ha) hs f + +lemma sup'_div₀ [ZeroLEOneClass G₀] [MulPosStrictMono G₀] [MulPosReflectLE G₀] [PosMulReflectLT G₀] + (ha : 0 < a) (f : ι → G₀) (s : Finset ι) (hs) : s.sup' hs f / a = s.sup' hs fun i ↦ f i / a := + map_finset_sup' (OrderIso.divRight₀ _ ha) hs f + +end GroupWithZero + +lemma sup_div₀ [LinearOrderedCommGroupWithZero G₀] [OrderBot G₀] {a : G₀} (ha : 0 < a) + (s : Finset ι) (f : ι → G₀) : s.sup f / a = s.sup fun i ↦ f i / a := by + obtain rfl | hs := s.eq_empty_or_nonempty + · simp [← show (0 : G₀) = ⊥ from bot_unique zero_le'] + rw [← Finset.sup'_eq_sup hs, ← Finset.sup'_eq_sup hs, sup'_div₀ (ha := ha)] + +end Finset diff --git a/Mathlib/Algebra/Order/Ring/Finset.lean b/Mathlib/Algebra/Order/Ring/Finset.lean deleted file mode 100644 index dd490e170eff4..0000000000000 --- a/Mathlib/Algebra/Order/Ring/Finset.lean +++ /dev/null @@ -1,34 +0,0 @@ -/- -Copyright (c) 2022 Eric Wieser. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Eric Wieser --/ -import Mathlib.Algebra.Order.Ring.Defs -import Mathlib.Data.Finset.Lattice.Fold - -/-! -# Algebraic properties of finitary supremum --/ - -namespace Finset -variable {ι R : Type*} [LinearOrderedSemiring R] {a b : ι → R} - -theorem sup_mul_le_mul_sup_of_nonneg [OrderBot R] (s : Finset ι) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup (a * b) ≤ s.sup a * s.sup b := - Finset.sup_le fun _i hi ↦ - mul_le_mul (le_sup hi) (le_sup hi) (hb _ hi) ((ha _ hi).trans <| le_sup hi) - -theorem mul_inf_le_inf_mul_of_nonneg [OrderTop R] (s : Finset ι) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf a * s.inf b ≤ s.inf (a * b) := - Finset.le_inf fun i hi ↦ mul_le_mul (inf_le hi) (inf_le hi) (Finset.le_inf hb) (ha i hi) - -theorem sup'_mul_le_mul_sup'_of_nonneg (s : Finset ι) (H : s.Nonempty) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.sup' H (a * b) ≤ s.sup' H a * s.sup' H b := - (sup'_le _ _) fun _i hi ↦ - mul_le_mul (le_sup' _ hi) (le_sup' _ hi) (hb _ hi) ((ha _ hi).trans <| le_sup' _ hi) - -theorem inf'_mul_le_mul_inf'_of_nonneg (s : Finset ι) (H : s.Nonempty) (ha : ∀ i ∈ s, 0 ≤ a i) - (hb : ∀ i ∈ s, 0 ≤ b i) : s.inf' H a * s.inf' H b ≤ s.inf' H (a * b) := - (le_inf' _ _) fun _i hi ↦ mul_le_mul (inf'_le _ hi) (inf'_le _ hi) (le_inf' _ _ hb) (ha _ hi) - -end Finset diff --git a/Mathlib/Analysis/Normed/Field/Lemmas.lean b/Mathlib/Analysis/Normed/Field/Lemmas.lean index 88affc76dbe77..2943b39e7e097 100644 --- a/Mathlib/Analysis/Normed/Field/Lemmas.lean +++ b/Mathlib/Analysis/Normed/Field/Lemmas.lean @@ -6,7 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl import Mathlib.Algebra.Group.AddChar import Mathlib.Algebra.Group.TypeTags.Finite -import Mathlib.Algebra.Order.Ring.Finset +import Mathlib.Algebra.Order.GroupWithZero.Finset import Mathlib.Analysis.Normed.Field.Basic import Mathlib.Analysis.Normed.Group.Bounded import Mathlib.Analysis.Normed.Group.Rat @@ -60,7 +60,7 @@ instance Pi.nonUnitalSeminormedRing {π : ι → Type*} [Fintype ι] Finset.univ.sup ((fun i => ‖x i‖₊) * fun i => ‖y i‖₊) := Finset.sup_mono_fun fun _ _ => norm_mul_le _ _ _ ≤ (Finset.univ.sup fun i => ‖x i‖₊) * Finset.univ.sup fun i => ‖y i‖₊ := - Finset.sup_mul_le_mul_sup_of_nonneg _ (fun _ _ => zero_le _) fun _ _ => zero_le _ + Finset.sup_mul_le_mul_sup_of_nonneg (fun _ _ => zero_le _) fun _ _ => zero_le _ } end NonUnitalSeminormedRing From dee58f263c603ed2915cb87fd3051be9ce27ce13 Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Mon, 25 Nov 2024 17:36:47 +0000 Subject: [PATCH 155/250] feat(NumberTheory/LSeries/PrimesInAP): proof of Dirichlet's Theorem (#19421) Co-authored-by: Johan Commelin --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 257 ++++++++++++++++++- docs/100.yaml | 2 + 2 files changed, 250 insertions(+), 9 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 6a882afa27f44..3cd0c083dd5d6 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Stoll -/ import Mathlib.NumberTheory.DirichletCharacter.Orthogonality -import Mathlib.NumberTheory.LSeries.DirichletContinuation import Mathlib.NumberTheory.LSeries.Linearity +import Mathlib.NumberTheory.LSeries.Nonvanishing import Mathlib.RingTheory.RootsOfUnity.AlgebraicallyClosed /-! @@ -15,12 +15,55 @@ The goal of this file is to prove **Dirichlet's Theorem**: If `q` is a positive and `a : ZMod q` is invertible, then there are infinitely many prime numbers `p` such that `(p : ZMod q) = a`. -This will be done in the following steps: - -1. Some auxiliary lemmas on infinite products and sums -2. Results on the von Mangoldt function restricted to a residue class -3. (TODO) Results on its L-series and an auxiliary function related to it -4. (TODO) Derivation of Dirichlet's Theorem +The main steps of the proof are as follows. +1. Define `ArithmeticFunction.vonMangoldt.residueClass a` for `a : ZMod q`, which is + a function `ℕ → ℝ` taking the value zero when `(n : ℤMod q) ≠ a` and `Λ n` else + (where `Λ` is the von Mangoldt function `ArithmeticFunction.vonMangoldt`; we have + `Λ (p^k) = log p` for prime powers and `Λ n = 0` otherwise.) +2. Show that this function can be written as a linear combination of functions + of the form `χ * Λ` (pointwise product) with Dirichlet characters `χ` mod `q`. + See `ArithmeticFunction.vonMangoldt.residueClass_eq`. +3. This implies that the L-series of `ArithmeticFunction.vonMangoldt.residueClass a` + agrees (on `re s > 1`) with the corresponding linear combination of negative logarithmic + derivatives of Dirichlet L-functions. + See `ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq`. +4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux a` that is + this linear combination of negative logarithmic derivatives of L-functions minus + `(q.totient)⁻¹/(s-1)`, which cancels the pole at `s = 1`. + See `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux` for the statement + that the auxiliary function agrees with the L-series of + `ArithmeticFunction.vonMangoldt.residueClass` up to the term `(q.totient)⁻¹/(s-1)`. +5. Show that the auxiliary function is continuous on `re s ≥ 1`; + see `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux`. + This relies heavily on the non-vanishing of Dirichlet L-functions on the *closed* + half-plane `re s ≥ 1` (`DirichletCharacter.LFunction_ne_zero_of_one_le_re`), which + in turn can only be stated since we know that the L-series of a Dirichlet character + extends to an entire function (unless the character is trivial; then there is a + simple pole at `s = 1`); see `DirichletCharacter.LFunction_eq_LSeries` + (contributed by David Loeffler). +6. Show that the sum of `Λ n / n` over any residue class, but *excluding* the primes, converges. + See `ArithmeticFunction.vonMangoldt.summable_residueClass_non_primes_div`. +7. Combining these ingredients, we can deduce that the sum of `Λ n / n` over + the *primes* in a residue class must diverge. + See `ArithmeticFunction.vonMangoldt.not_summable_residueClass_prime_div`. +8. This finally easily implies that there must be infinitely many primes in the residue class. + +## Definitions + +* `ArithmeticFunction.vonMangoldt.residueClass a` (see above). +* `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux` (see above). + +## Main Result + +We give two versions of **Dirichlet's Theorem**: +* `Nat.setOf_prime_and_eq_mod_infinite` states that the set of primes `p` + such that `(p : ZMod q) = a` is infinite (when `a` is invertible in `ZMod q`). +* `Nat.forall_exists_prime_gt_and_eq_mod` states that for any natural number `n` + there is a prime `p > n` such that `(p : ZMod q) = a`. + +## Tags + +prime number, arithmetic progression, residue class, Dirichlet's Theorem -/ /-! @@ -214,7 +257,7 @@ lemma summable_residueClass_non_primes_div : variable [NeZero q] {a} /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function with Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet charaters. -/ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : residueClass a n = (q.totient : ℂ)⁻¹ * ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ * χ n * vonMangoldt n := by @@ -224,7 +267,7 @@ lemma residueClass_apply (ha : IsUnit a) (n : ℕ) : ite_mul, zero_mul, ↓reduceIte, ite_self] /-- We can express `ArithmeticFunction.vonMangoldt.residueClass` as a linear combination -of twists of the von Mangoldt function with Dirichlet charaters. -/ +of twists of the von Mangoldt function by Dirichlet charaters. -/ lemma residueClass_eq (ha : IsUnit a) : ↗(residueClass a) = (q.totient : ℂ)⁻¹ • ∑ χ : DirichletCharacter ℂ q, χ a⁻¹ • (fun n : ℕ ↦ χ n * vonMangoldt n) := by @@ -248,6 +291,202 @@ lemma LSeries_residueClass_eq (ha : IsUnit a) {s : ℂ} (hs : 1 < s.re) : simp only [Pi.smul_apply, residueClass_apply ha, smul_eq_mul, ← mul_assoc, mul_inv_cancel_of_invertible, one_mul, Finset.sum_apply, Pi.mul_apply] +variable (a) + +open Classical in +/-- The auxiliary function used, e.g., with the Wiener-Ikehara Theorem to prove +Dirichlet's Theorem. On `re s > 1`, it agrees with the L-series of the von Mangoldt +function restricted to the residue class `a : ZMod q` minus the principal part +`(q.totient)⁻¹/(s-1)` of the pole at `s = 1`; +see `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux`. -/ +noncomputable +abbrev LfunctionResidueClassAux (s : ℂ) : ℂ := + (q.totient : ℂ)⁻¹ * (-deriv (LFunctionTrivChar₁ q) s / LFunctionTrivChar₁ q s - + ∑ χ ∈ ({1}ᶜ : Finset (DirichletCharacter ℂ q)), χ a⁻¹ * deriv (LFunction χ) s / LFunction χ s) + +/-- The auxiliary function is continuous away from the zeros of the L-functions of the Dirichlet +characters mod `q` (including at `s = 1`). -/ +lemma continuousOn_LfunctionResidueClassAux' : + ContinuousOn (LfunctionResidueClassAux a) + {s | s = 1 ∨ ∀ χ : DirichletCharacter ℂ q, LFunction χ s ≠ 0} := by + rw [show LfunctionResidueClassAux a = fun s ↦ _ from rfl] + simp only [LfunctionResidueClassAux, sub_eq_add_neg] + refine continuousOn_const.mul <| ContinuousOn.add ?_ ?_ + · refine (continuousOn_neg_logDeriv_LFunctionTrivChar₁ q).mono fun s hs ↦ ?_ + have := LFunction_ne_zero_of_one_le_re (1 : DirichletCharacter ℂ q) (s := s) + simp only [ne_eq, Set.mem_setOf_eq] at hs + tauto + · simp only [← Finset.sum_neg_distrib, mul_div_assoc, ← mul_neg, ← neg_div] + refine continuousOn_finset_sum _ fun χ hχ ↦ continuousOn_const.mul ?_ + replace hχ : χ ≠ 1 := by simpa only [ne_eq, Finset.mem_compl, Finset.mem_singleton] using hχ + refine (continuousOn_neg_logDeriv_LFunction_of_nontriv hχ).mono fun s hs ↦ ?_ + simp only [ne_eq, Set.mem_setOf_eq] at hs + rcases hs with rfl | hs + · simp only [ne_eq, Set.mem_setOf_eq, one_re, le_refl, + LFunction_ne_zero_of_one_le_re χ (.inl hχ), not_false_eq_true] + · exact hs χ + +/-- The L-series of the von Mangoldt function restricted to the prime residue class `a` mod `q` +is continuous on `re s ≥ 1` except for a simple pole at `s = 1` with residue `(q.totient)⁻¹`. +The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux` +is equivalent. -/ +lemma continuousOn_LfunctionResidueClassAux : + ContinuousOn (LfunctionResidueClassAux a) {s | 1 ≤ s.re} := by + refine (continuousOn_LfunctionResidueClassAux' a).mono fun s hs ↦ ?_ + rcases eq_or_ne s 1 with rfl | hs₁ + · simp only [ne_eq, Set.mem_setOf_eq, true_or] + · simp only [ne_eq, Set.mem_setOf_eq, hs₁, false_or] + exact fun χ ↦ LFunction_ne_zero_of_one_le_re χ (.inr hs₁) <| Set.mem_setOf.mp hs + +variable {a} + +open scoped LSeries.notation + +/-- The auxiliary function agrees on `re s > 1` with the L-series of the von Mangoldt function +restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` +of its pole at `s = 1`. -/ +lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : + Set.EqOn (LfunctionResidueClassAux a) + (fun s ↦ L ↗(residueClass a) s - (q.totient : ℂ)⁻¹ / (s - 1)) + {s | 1 < s.re} := by + intro s hs + replace hs := Set.mem_setOf.mp hs + simp only [LSeries_residueClass_eq ha hs, LfunctionResidueClassAux] + rw [neg_div, ← neg_add', mul_neg, ← neg_mul, div_eq_mul_one_div (q.totient : ℂ)⁻¹, + sub_eq_add_neg, ← neg_mul, ← mul_add] + congrm (_ * ?_) + -- this should be easier, but `IsUnit.inv ha` does not work here + have ha' : IsUnit a⁻¹ := isUnit_of_dvd_one ⟨a, (ZMod.inv_mul_of_unit a ha).symm⟩ + classical -- for `Fintype.sum_eq_add_sum_compl` + rw [Fintype.sum_eq_add_sum_compl 1, MulChar.one_apply ha', one_mul, add_right_comm] + simp only [mul_div_assoc] + congrm (?_ + _) + have hs₁ : s ≠ 1 := fun h ↦ ((h ▸ hs).trans_eq one_re).false + rw [deriv_LFunctionTrivChar₁_apply_of_ne_one _ hs₁, LFunctionTrivChar₁, + Function.update_noteq hs₁, LFunctionTrivChar, add_div, + mul_div_mul_left _ _ (sub_ne_zero_of_ne hs₁)] + conv_lhs => enter [2, 1]; rw [← mul_one (LFunction ..)] + rw [mul_comm _ 1, mul_div_mul_right _ _ <| LFunction_ne_zero_of_one_le_re 1 (.inr hs₁) hs.le] + +/-- The auxiliary function takes real values for real arguments `x > 1`. -/ +lemma LfunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : + LfunctionResidueClassAux a x = (LfunctionResidueClassAux a x).re := by + rw [eqOn_LfunctionResidueClassAux ha hx] + simp only [sub_re, ofReal_sub] + congr 1 + · rw [LSeries, re_tsum <| LSeriesSummable_of_abscissaOfAbsConv_lt_re <| + (abscissaOfAbsConv_residueClass_le_one a).trans_lt <| by norm_cast] + push_cast + refine tsum_congr fun n ↦ ?_ + rcases eq_or_ne n 0 with rfl | hn + · simp only [term_zero, zero_re, ofReal_zero] + · simp only [term_of_ne_zero hn, ← ofReal_natCast n, ← ofReal_cpow n.cast_nonneg, ← ofReal_div, + ofReal_re] + · rw [show (q.totient : ℂ) = (q.totient : ℝ) from rfl, ← ofReal_one, ← ofReal_sub, ← ofReal_inv, + ← ofReal_div, ofReal_re] + +variable {q : ℕ} [NeZero q] {a : ZMod q} + +/-- As `x` approaches `1` from the right along the real axis, the L-series of +`ArithmeticFunction.vonMangoldt.residueClass` is bounded below by `(q.totient)⁻¹/(x-1) - C`. -/ +lemma LSeries_residueClass_lower_bound (ha : IsUnit a) : + ∃ C : ℝ, ∀ {x : ℝ} (_ : x ∈ Set.Ioc 1 2), + (q.totient : ℝ)⁻¹ / (x - 1) - C ≤ ∑' n, residueClass a n / (n : ℝ) ^ x := by + have H {x : ℝ} (hx : 1 < x) : + ∑' n, residueClass a n / (n : ℝ) ^ x = + (LfunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by + refine ofReal_injective ?_ + simp only [ofReal_tsum, ofReal_div, ofReal_cpow (Nat.cast_nonneg _), ofReal_natCast, + ofReal_add, ofReal_inv, ofReal_sub, ofReal_one] + simp_rw [← LfunctionResidueClassAux_real ha hx, + eqOn_LfunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, + LSeries, term] + refine tsum_congr fun n ↦ ?_ + split_ifs with hn + · simp only [hn, residueClass_apply_zero, ofReal_zero, zero_div] + · rfl + have : ContinuousOn (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) (Set.Icc 1 2) := + continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LfunctionResidueClassAux a) + (fun ⦃x⦄ a ↦ trivial) |>.comp continuous_ofReal.continuousOn fun x hx ↦ by + simpa only [Set.mem_setOf_eq, ofReal_re] using hx.1 + obtain ⟨C, hC⟩ := bddBelow_def.mp <| IsCompact.bddBelow_image isCompact_Icc this + replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LfunctionResidueClassAux a x).re := + hC (LfunctionResidueClassAux a x).re <| + Set.mem_image_of_mem (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) hx + refine ⟨-C, fun {x} hx ↦ ?_⟩ + rw [H hx.1, add_comm, sub_neg_eq_add, add_le_add_iff_left] + exact hC <| Set.mem_Icc_of_Ioc hx + +open vonMangoldt Filter Topology in +/-- The function `n ↦ Λ n / n` restricted to primes in an invertible residue class +is not summable. This then implies that there must be infinitely many such primes. -/ +lemma not_summable_residueClass_prime_div (ha : IsUnit a) : + ¬ Summable fun n : ℕ ↦ (if n.Prime then residueClass a n else 0) / n := by + intro H + have key : Summable fun n : ℕ ↦ residueClass a n / n := by + convert (summable_residueClass_non_primes_div a).add H using 2 with n + simp only [← add_div, ite_add_ite, zero_add, add_zero, ite_self] + let C := ∑' n, residueClass a n / n + have H₁ {x : ℝ} (hx : 1 < x) : ∑' n, residueClass a n / (n : ℝ) ^ x ≤ C := by + refine tsum_le_tsum (fun n ↦ ?_) ?_ key + · rcases n.eq_zero_or_pos with rfl | hn + · simp only [Nat.cast_zero, Real.zero_rpow (zero_lt_one.trans hx).ne', div_zero, le_refl] + · refine div_le_div_of_nonneg_left (residueClass_nonneg a _) (mod_cast hn) ?_ + conv_lhs => rw [← Real.rpow_one n] + exact Real.rpow_le_rpow_of_exponent_le (by norm_cast) hx.le + · exact summable_real_of_abscissaOfAbsConv_lt <| + (abscissaOfAbsConv_residueClass_le_one a).trans_lt <| mod_cast hx + obtain ⟨C', hC'⟩ := LSeries_residueClass_lower_bound ha + have H₁ {x} (hx : x ∈ Set.Ioc 1 2) : (q.totient : ℝ)⁻¹ ≤ (C + C') * (x - 1) := + (div_le_iff₀ <| sub_pos.mpr hx.1).mp <| + sub_le_iff_le_add.mp <| (hC' hx).trans (H₁ hx.1) + have hq : 0 < (q.totient : ℝ)⁻¹ := inv_pos.mpr (mod_cast q.totient.pos_of_neZero) + rcases le_or_lt (C + C') 0 with h₀ | h₀ + · have := hq.trans_le (H₁ (Set.right_mem_Ioc.mpr one_lt_two)) + rw [show (2 : ℝ) - 1 = 1 by norm_num, mul_one] at this + exact (this.trans_le h₀).false + · obtain ⟨ξ, hξ₁, hξ₂⟩ : ∃ ξ ∈ Set.Ioc 1 2, (C + C') * (ξ - 1) < (q.totient : ℝ)⁻¹ := by + refine ⟨min (1 + (q.totient : ℝ)⁻¹ / (C + C') / 2) 2, ⟨?_, min_le_right ..⟩, ?_⟩ + · simpa only [lt_inf_iff, lt_add_iff_pos_right, Nat.ofNat_pos, div_pos_iff_of_pos_right, + Nat.one_lt_ofNat, and_true] using div_pos hq h₀ + · rw [← min_sub_sub_right, add_sub_cancel_left, ← lt_div_iff₀' h₀] + exact (min_le_left ..).trans_lt <| div_lt_self (div_pos hq h₀) one_lt_two + exact ((H₁ hξ₁).trans_lt hξ₂).false + end ArithmeticFunction.vonMangoldt end arith_prog + +/-! +### Dirichlet's Theorem +-/ + +section DirichletsTheorem + +namespace Nat + +open ArithmeticFunction vonMangoldt + +variable {q : ℕ} [NeZero q] {a : ZMod q} + +/-- **Dirichlet's Theorem** on primes in arithmetic progression: if `q` is a positive +integer and `a : ZMod q` is a unit, then there are infintely many prime numbers `p` +such that `(p : ZMod q) = a`. -/ +theorem setOf_prime_and_eq_mod_infinite (ha : IsUnit a) : + {p : ℕ | p.Prime ∧ (p : ZMod q) = a}.Infinite := by + by_contra H + rw [Set.not_infinite] at H + exact not_summable_residueClass_prime_div ha <| + summable_of_finite_support <| support_residueClass_prime_div a ▸ H + +/-- **Dirichlet's Theorem** on primes in arithmetic progression: if `q` is a positive +integer and `a : ZMod q` is a unit, then there are infintely many prime numbers `p` +such that `(p : ZMod q) = a`. -/ +theorem forall_exists_prime_gt_and_eq_mod (ha : IsUnit a) (n : ℕ) : + ∃ p > n, p.Prime ∧ (p : ZMod q) = a := by + obtain ⟨p, hp₁, hp₂⟩ := Set.infinite_iff_exists_gt.mp (setOf_prime_and_eq_mod_infinite ha) n + exact ⟨p, hp₂.gt, Set.mem_setOf.mp hp₁⟩ + +end Nat + +end DirichletsTheorem diff --git a/docs/100.yaml b/docs/100.yaml index 31f5837e872ca..550cf7682b392 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -182,6 +182,8 @@ title : The Central Limit Theorem 48: title : Dirichlet’s Theorem + decl : Nat.setOf_prime_and_eq_mod_infinite + author : David Loeffler, Michael Stoll 49: title : The Cayley-Hamilton Theorem decl : Matrix.aeval_self_charpoly From 87489d6ce898fbcf567404c1f01e98f63bb1732f Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Mon, 25 Nov 2024 17:36:48 +0000 Subject: [PATCH 156/250] chore(AlgebraicTopology): create Quasicategory directory (#19463) Moves: - AlgebraicTopology.SimplicialSet.Quasicategory -> AlgebraicTopology.Quasicategory.Basic Co-Authored-By: [Johan Commelin](https://github.com/jcommelin) --- Mathlib.lean | 2 +- .../Quasicategory.lean => Quasicategory/Basic.lean} | 7 ++++--- Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean | 2 +- Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean | 2 +- 4 files changed, 7 insertions(+), 6 deletions(-) rename Mathlib/AlgebraicTopology/{SimplicialSet/Quasicategory.lean => Quasicategory/Basic.lean} (91%) diff --git a/Mathlib.lean b/Mathlib.lean index 7206497b94d35..5785a8dab0524 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1034,6 +1034,7 @@ import Mathlib.AlgebraicTopology.FundamentalGroupoid.PUnit import Mathlib.AlgebraicTopology.FundamentalGroupoid.Product import Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected import Mathlib.AlgebraicTopology.MooreComplex +import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject @@ -1043,7 +1044,6 @@ import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path -import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal import Mathlib.AlgebraicTopology.SingularSet import Mathlib.AlgebraicTopology.SplitSimplicialObject diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean similarity index 91% rename from Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean rename to Mathlib/AlgebraicTopology/Quasicategory/Basic.lean index 381e65660f632..55b27d884364c 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Quasicategory.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean @@ -13,13 +13,14 @@ In this file we define quasicategories, a common model of infinity categories. We show that every Kan complex is a quasicategory. -In `Mathlib/AlgebraicTopology/Nerve.lean` -we show (TODO) that the nerve of a category is a quasicategory. +In `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean` +we show that the nerve of a category is a quasicategory. ## TODO - Generalize the definition to higher universes. - See the corresponding TODO in `Mathlib/AlgebraicTopology/KanComplex.lean`. + See the corresponding TODO in + `Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean`. -/ diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean index 531bc99ebc39c..87cf0ca7bf9e5 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/KanComplex.lean @@ -10,7 +10,7 @@ import Mathlib.AlgebraicTopology.SimplicialSet.Basic # Kan complexes In this file we give the definition of Kan complexes. -In `Mathlib/AlgebraicTopology/Quasicategory.lean` +In `Mathlib/AlgebraicTopology/Quasicategory/Basic.lean` we show that every Kan complex is a quasicategory. ## TODO diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index ca5432b7bbe6b..cef9a156cab5a 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -3,9 +3,9 @@ Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ +import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path -import Mathlib.AlgebraicTopology.SimplicialSet.Quasicategory import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction import Mathlib.CategoryTheory.Functor.KanExtension.Basic From f44426d3cc139d334553648bdd614cf74ca02d79 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:36:50 +0000 Subject: [PATCH 157/250] feat(FreeGroup): some basic lemmas (#19468) From LeanCamCombi --- Mathlib/GroupTheory/FreeGroup/Basic.lean | 16 ++++++++++++++-- 1 file changed, 14 insertions(+), 2 deletions(-) diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index d365be175ae24..2c47200b1ac4d 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -48,6 +48,7 @@ free group, Newman's diamond lemma, Church-Rosser theorem -/ open Relation +open scoped List universe u v w @@ -906,6 +907,9 @@ def reduce : (L : List (α × Bool)) -> List (α × Bool) := List.casesOn ih [hd1] fun hd2 tl2 => if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 +@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl +@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl + @[to_additive (attr := simp)] theorem reduce.cons (x) : reduce (x :: L) = @@ -1099,11 +1103,19 @@ theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (redu have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev rwa [invRev_invRev] at this -@[to_additive] -theorem toWord_inv {x : FreeGroup α} : x⁻¹.toWord = invRev x.toWord := by +@[to_additive (attr := simp)] +theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by rcases x with ⟨L⟩ rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] +@[to_additive] +lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by + refine Red.sublist ?_ + have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by + rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] + rw [this] + exact FreeGroup.reduce.red + /-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ @[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := From ebf44be306e1f48663536f8764e535404630d955 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Mon, 25 Nov 2024 17:43:12 +0000 Subject: [PATCH 158/250] Trigger CI for https://github.com/leanprover/lean4/pull/6189 --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 67a07f2e7db05..ddb7b5fad8fca 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0e076cbda9033e5cde14aa444346c6a191f33972", + "rev": "e29a2ac4727c46b72b4f7bb971e35fea681ef878", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6189", From b98172a9b91d3ffb0c8ea2f82646639724f8a3d4 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 17:50:07 +0000 Subject: [PATCH 159/250] feat: `Icc a b + Icc c d = Icc (a + c) (b + d)` (#19398) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From GrowthInGroups Moves: * `inv_Ioi` -> `inv_Ioi₀` --- .../Order/Group/Pointwise/Interval.lean | 114 ++++++++++++------ Mathlib/Analysis/Calculus/LHopital.lean | 12 +- Mathlib/Analysis/Convex/Topology.lean | 4 +- .../Gaussian/GaussianIntegral.lean | 2 +- .../MeasureTheory/Measure/Haar/OfBasis.lean | 2 +- .../Measure/Lebesgue/Integral.lean | 4 +- Mathlib/Topology/Algebra/Order/Field.lean | 2 +- 7 files changed, 87 insertions(+), 53 deletions(-) diff --git a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean index e6846ddbc75b3..fb5b4cae02b8a 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/Interval.lean @@ -114,6 +114,67 @@ theorem Ici_mul_Ioi_subset' (a b : α) : Ici a * Ioi b ⊆ Ioi (a * b) := by end ContravariantLT +section LinearOrderedCommMonoid +variable [LinearOrderedCommMonoid α] [MulLeftReflectLE α] [ExistsMulOfLE α] {a b c d : α} + +-- TODO: Generalise to arbitrary actions using a `smul` version of `MulLeftMono` +@[to_additive (attr := simp)] +lemma smul_Icc (a b c : α) : a • Icc b c = Icc (a * b) (a * c) := by + ext x + constructor + · rintro ⟨y, ⟨hby, hyc⟩, rfl⟩ + exact ⟨mul_le_mul_left' hby _, mul_le_mul_left' hyc _⟩ + · rintro ⟨habx, hxac⟩ + obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le habx + refine ⟨b * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, (mul_assoc ..).symm⟩ + rwa [mul_assoc, mul_le_mul_iff_left] at hxac + +@[to_additive] +lemma Icc_mul_Icc (hab : a ≤ b) (hcd : c ≤ d) : Icc a b * Icc c d = Icc (a * c) (b * d) := by + refine (Icc_mul_Icc_subset' _ _ _ _).antisymm fun x ⟨hacx, hxbd⟩ ↦ ?_ + obtain hxbc | hbcx := le_total x (b * c) + · obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le hacx + refine ⟨a * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, c, left_mem_Icc.2 hcd, mul_right_comm ..⟩ + rwa [mul_right_comm, mul_le_mul_iff_right] at hxbc + · obtain ⟨y, hy, rfl⟩ := exists_one_le_mul_of_le hbcx + refine ⟨b, right_mem_Icc.2 hab, c * y, ⟨le_mul_of_one_le_right' hy, ?_⟩, (mul_assoc ..).symm⟩ + rwa [mul_assoc, mul_le_mul_iff_left] at hxbd + +end LinearOrderedCommMonoid + +section OrderedCommGroup +variable [OrderedCommGroup α] + +@[to_additive (attr := simp)] lemma inv_Ici (a : α) : (Ici a)⁻¹ = Iic a⁻¹ := ext fun _x ↦ le_inv' +@[to_additive (attr := simp)] lemma inv_Iic (a : α) : (Iic a)⁻¹ = Ici a⁻¹ := ext fun _x ↦ inv_le' +@[to_additive (attr := simp)] lemma inv_Ioi (a : α) : (Ioi a)⁻¹ = Iio a⁻¹ := ext fun _x ↦ lt_inv' +@[to_additive (attr := simp)] lemma inv_Iio (a : α) : (Iio a)⁻¹ = Ioi a⁻¹ := ext fun _x ↦ inv_lt' + +@[to_additive (attr := simp)] +lemma inv_Icc (a b : α) : (Icc a b)⁻¹ = Icc b⁻¹ a⁻¹ := by simp [← Ici_inter_Iic, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ico (a b : α) : (Ico a b)⁻¹ = Ioc b⁻¹ a⁻¹ := by + simp [← Ici_inter_Iio, ← Ioi_inter_Iic, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ioc (a b : α) : (Ioc a b)⁻¹ = Ico b⁻¹ a⁻¹ := by + simp [← Ioi_inter_Iic, ← Ici_inter_Iio, inter_comm] + +@[to_additive (attr := simp)] +lemma inv_Ioo (a b : α) : (Ioo a b)⁻¹ = Ioo b⁻¹ a⁻¹ := by simp [← Ioi_inter_Iio, inter_comm] + +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ici := neg_Ici +@[deprecated (since := "2024-11-23")] alias preimage_neg_Iic := neg_Iic +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioi := neg_Ioi +@[deprecated (since := "2024-11-23")] alias preimage_neg_Iio := neg_Iio +@[deprecated (since := "2024-11-23")] alias preimage_neg_Icc := neg_Icc +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ico := neg_Ico +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioc := neg_Ioc +@[deprecated (since := "2024-11-23")] alias preimage_neg_Ioo := neg_Ioo + +end OrderedCommGroup + section OrderedAddCommGroup variable [OrderedAddCommGroup α] (a b c : α) @@ -192,41 +253,6 @@ theorem preimage_add_const_Ioc : (fun x => x + a) ⁻¹' Ioc b c = Ioc (b - a) ( theorem preimage_add_const_Ioo : (fun x => x + a) ⁻¹' Ioo b c = Ioo (b - a) (c - a) := by simp [← Ioi_inter_Iio] -/-! -### Preimages under `x ↦ -x` --/ - - -@[simp] -theorem preimage_neg_Ici : -Ici a = Iic (-a) := - ext fun _x => le_neg - -@[simp] -theorem preimage_neg_Iic : -Iic a = Ici (-a) := - ext fun _x => neg_le - -@[simp] -theorem preimage_neg_Ioi : -Ioi a = Iio (-a) := - ext fun _x => lt_neg - -@[simp] -theorem preimage_neg_Iio : -Iio a = Ioi (-a) := - ext fun _x => neg_lt - -@[simp] -theorem preimage_neg_Icc : -Icc a b = Icc (-b) (-a) := by simp [← Ici_inter_Iic, inter_comm] - -@[simp] -theorem preimage_neg_Ico : -Ico a b = Ioc (-b) (-a) := by - simp [← Ici_inter_Iio, ← Ioi_inter_Iic, inter_comm] - -@[simp] -theorem preimage_neg_Ioc : -Ioc a b = Ico (-b) (-a) := by - simp [← Ioi_inter_Iic, ← Ici_inter_Iio, inter_comm] - -@[simp] -theorem preimage_neg_Ioo : -Ioo a b = Ioo (-b) (-a) := by simp [← Ioi_inter_Iio, inter_comm] - /-! ### Preimages under `x ↦ x - a` -/ @@ -433,6 +459,15 @@ theorem Iio_add_bij : BijOn (· + a) (Iio b) (Iio (b + a)) := end OrderedAddCommGroup +section LinearOrderedCommGroup +variable [LinearOrderedCommGroup α] + +@[to_additive (attr := simp)] +lemma inv_uIcc (a b : α) : [[a, b]]⁻¹ = [[a⁻¹, b⁻¹]] := by + simp only [uIcc, inv_Icc, inv_sup, inv_inf] + +end LinearOrderedCommGroup + section LinearOrderedAddCommGroup variable [LinearOrderedAddCommGroup α] (a b c d : α) @@ -445,10 +480,9 @@ theorem preimage_const_add_uIcc : (fun x => a + x) ⁻¹' [[b, c]] = [[b - a, c theorem preimage_add_const_uIcc : (fun x => x + a) ⁻¹' [[b, c]] = [[b - a, c - a]] := by simpa only [add_comm] using preimage_const_add_uIcc a b c --- TODO: Why is the notation `-[[a, b]]` broken? -@[simp] -theorem preimage_neg_uIcc : @Neg.neg (Set α) Set.neg [[a, b]] = [[-a, -b]] := by - simp only [← Icc_min_max, preimage_neg_Icc, min_neg_neg, max_neg_neg] +@[deprecated neg_uIcc (since := "2024-11-23")] +theorem preimage_neg_uIcc : -[[a, b]] = [[-a, -b]] := by + simp only [← Icc_min_max, neg_Icc, min_neg_neg, max_neg_neg] @[simp] theorem preimage_sub_const_uIcc : (fun x => x - a) ⁻¹' [[b, c]] = [[b + a, c + a]] := by @@ -750,7 +784,7 @@ theorem inv_Ioo_0_left {a : α} (ha : 0 < a) : (Ioo 0 a)⁻¹ = Ioi a⁻¹ := by inv_inv a ▸ (inv_lt_inv₀ ((inv_pos.2 ha).trans h) (inv_pos.2 ha)).2 h⟩⟩ -theorem inv_Ioi {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by +theorem inv_Ioi₀ {a : α} (ha : 0 < a) : (Ioi a)⁻¹ = Ioo 0 a⁻¹ := by rw [inv_eq_iff_eq_inv, inv_Ioo_0_left (inv_pos.2 ha), inv_inv] theorem image_const_mul_Ioi_zero {k : Type*} [LinearOrderedField k] {x : k} (hx : 0 < x) : diff --git a/Mathlib/Analysis/Calculus/LHopital.lean b/Mathlib/Analysis/Calculus/LHopital.lean index f1e691d8342cf..66b0aac6e04cc 100644 --- a/Mathlib/Analysis/Calculus/LHopital.lean +++ b/Mathlib/Analysis/Calculus/LHopital.lean @@ -110,11 +110,11 @@ theorem lhopital_zero_left_on_Ioo (hab : a < b) (hff' : ∀ x ∈ Ioo a b, HasDe comp x (hff' (-x) hx) (hasDerivAt_neg x) have hdng : ∀ x ∈ -Ioo a b, HasDerivAt (g ∘ Neg.neg) (g' (-x) * -1) x := fun x hx => comp x (hgg' (-x) hx) (hasDerivAt_neg x) - rw [preimage_neg_Ioo] at hdnf - rw [preimage_neg_Ioo] at hdng + rw [neg_Ioo] at hdnf + rw [neg_Ioo] at hdng have := lhopital_zero_right_on_Ioo (neg_lt_neg hab) hdnf hdng (by intro x hx h - apply hg' _ (by rw [← preimage_neg_Ioo] at hx; exact hx) + apply hg' _ (by rw [← neg_Ioo] at hx; exact hx) rwa [mul_comm, ← neg_eq_neg_one_mul, neg_eq_zero] at h) (hfb.comp tendsto_neg_nhdsWithin_Ioi_neg) (hgb.comp tendsto_neg_nhdsWithin_Ioi_neg) (by @@ -176,12 +176,12 @@ theorem lhopital_zero_atBot_on_Iio (hff' : ∀ x ∈ Iio a, HasDerivAt f (f' x) comp x (hff' (-x) hx) (hasDerivAt_neg x) have hdng : ∀ x ∈ -Iio a, HasDerivAt (g ∘ Neg.neg) (g' (-x) * -1) x := fun x hx => comp x (hgg' (-x) hx) (hasDerivAt_neg x) - rw [preimage_neg_Iio] at hdnf - rw [preimage_neg_Iio] at hdng + rw [neg_Iio] at hdnf + rw [neg_Iio] at hdng have := lhopital_zero_atTop_on_Ioi hdnf hdng (by intro x hx h - apply hg' _ (by rw [← preimage_neg_Iio] at hx; exact hx) + apply hg' _ (by rw [← neg_Iio] at hx; exact hx) rwa [mul_comm, ← neg_eq_neg_one_mul, neg_eq_zero] at h) (hfbot.comp tendsto_neg_atTop_atBot) (hgbot.comp tendsto_neg_atTop_atBot) (by diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index a38bd1386b8ea..86c3bddbec48e 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -360,8 +360,8 @@ theorem Convex.closure_subset_image_homothety_interior_of_one_lt {s : Set E} (hs refine ⟨homothety x t⁻¹ y, hs.openSegment_interior_closure_subset_interior hx hy ?_, (AffineEquiv.homothetyUnitsMulHom x (Units.mk0 t hne)).apply_symm_apply y⟩ - rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi (zero_lt_one' ℝ), ← image_inv, image_image, - homothety_eq_lineMap] + rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv, + image_image, homothety_eq_lineMap] exact mem_image_of_mem _ ht /-- If we dilate a convex set about a point in its interior by a scale `t > 1`, the interior of diff --git a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean index 75fc9f00bf0db..b88d8363dabf7 100644 --- a/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean +++ b/Mathlib/Analysis/SpecialFunctions/Gaussian/GaussianIntegral.lean @@ -109,7 +109,7 @@ theorem integrable_rpow_mul_exp_neg_mul_sq {b : ℝ} (hb : 0 < b) {s : ℝ} (hs refine ⟨?_, integrableOn_rpow_mul_exp_neg_mul_sq hb hs⟩ rw [← (Measure.measurePreserving_neg (volume : Measure ℝ)).integrableOn_comp_preimage (Homeomorph.neg ℝ).measurableEmbedding] - simp only [Function.comp_def, neg_sq, neg_preimage, preimage_neg_Iio, neg_neg, neg_zero] + simp only [Function.comp_def, neg_sq, neg_preimage, neg_Iio, neg_neg, neg_zero] apply Integrable.mono' (integrableOn_rpow_mul_exp_neg_mul_sq hb hs) · apply Measurable.aestronglyMeasurable exact (measurable_id'.neg.pow measurable_const).mul diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index 7f76b5cbe8048..a0db8d8e207b9 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -116,7 +116,7 @@ theorem parallelepiped_orthonormalBasis_one_dim (b : OrthonormalBasis ι ℝ ℝ · right simp_rw [H, parallelepiped, Algebra.id.smul_eq_mul, A] simp only [F, Finset.univ_unique, Fin.default_eq_zero, mul_neg, mul_one, Finset.sum_neg_distrib, - Finset.sum_singleton, ← image_comp, Function.comp, image_neg, preimage_neg_Icc, neg_zero] + Finset.sum_singleton, ← image_comp, Function.comp, image_neg, neg_Icc, neg_zero] theorem parallelepiped_eq_sum_segment (v : ι → E) : parallelepiped v = ∑ i, segment ℝ 0 (v i) := by ext diff --git a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean index e4968335130b1..e72be147a0f25 100644 --- a/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean +++ b/Mathlib/MeasureTheory/Measure/Lebesgue/Integral.lean @@ -83,7 +83,7 @@ theorem integral_comp_neg_Iic {E : Type*} [NormedAddCommGroup E] [NormedSpace (Homeomorph.neg ℝ).isClosedEmbedding.measurableEmbedding have := MeasurableEmbedding.setIntegral_map (μ := volume) A f (Ici (-c)) rw [Measure.map_neg_eq_self (volume : Measure ℝ)] at this - simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, preimage_neg_Ici, neg_neg] + simp_rw [← integral_Ici_eq_integral_Ioi, this, neg_preimage, neg_Ici, neg_neg] /- @[simp] Porting note: Linter complains it does not apply to itself. Although it does apply to itself, it does not apply when `f` is more complicated -/ @@ -102,7 +102,7 @@ theorem integral_comp_abs {f : ℝ → ℝ} : rw [← Measure.map_neg_eq_self (volume : Measure ℝ)] let m : MeasurableEmbedding fun x : ℝ => -x := (Homeomorph.neg ℝ).measurableEmbedding rw [m.integrableOn_map_iff] - simp_rw [Function.comp_def, abs_neg, neg_preimage, preimage_neg_Iic, neg_zero] + simp_rw [Function.comp_def, abs_neg, neg_preimage, neg_Iic, neg_zero] exact integrableOn_Ici_iff_integrableOn_Ioi.mpr hf calc _ = (∫ x in Iic 0, f |x|) + ∫ x in Ioi 0, f |x| := by diff --git a/Mathlib/Topology/Algebra/Order/Field.lean b/Mathlib/Topology/Algebra/Order/Field.lean index ee7ea3caa6442..29f1d4890fd66 100644 --- a/Mathlib/Topology/Algebra/Order/Field.lean +++ b/Mathlib/Topology/Algebra/Order/Field.lean @@ -114,7 +114,7 @@ theorem Filter.Tendsto.neg_mul_atBot {C : 𝕜} (hC : C < 0) (hf : Tendsto f l ( @[simp] lemma inv_atTop₀ : (atTop : Filter 𝕜)⁻¹ = 𝓝[>] 0 := (((atTop_basis_Ioi' (0 : 𝕜)).map _).comp_surjective inv_surjective).eq_of_same_basis <| - (nhdsWithin_Ioi_basis _).congr (by simp) fun a ha ↦ by simp [inv_Ioi (inv_pos.2 ha)] + (nhdsWithin_Ioi_basis _).congr (by simp) fun a ha ↦ by simp [inv_Ioi₀ (inv_pos.2 ha)] @[simp] lemma inv_nhdsWithin_Ioi_zero : (𝓝[>] (0 : 𝕜))⁻¹ = atTop := by rw [← inv_atTop₀, inv_inv] From 3646342693db92ccb97e120635ea7d068d60e738 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 18:11:57 +0000 Subject: [PATCH 160/250] feat(MvPolynomial): more lemmas about `finSuccEquiv` (#19201) Co-authored-by: Andrew Yang --- Mathlib/Algebra/MvPolynomial/Equiv.lean | 39 ++++++++++++++++++------- Mathlib/Data/Finsupp/Fin.lean | 5 ++++ 2 files changed, 33 insertions(+), 11 deletions(-) diff --git a/Mathlib/Algebra/MvPolynomial/Equiv.lean b/Mathlib/Algebra/MvPolynomial/Equiv.lean index 6b77addf42b8d..ebcd5678aa4a4 100644 --- a/Mathlib/Algebra/MvPolynomial/Equiv.lean +++ b/Mathlib/Algebra/MvPolynomial/Equiv.lean @@ -336,7 +336,6 @@ theorem finSuccEquiv_eq : rfl · refine Fin.cases ?_ ?_ i <;> simp [finSuccEquiv] -@[simp] theorem finSuccEquiv_apply (p : MvPolynomial (Fin (n + 1)) R) : finSuccEquiv R n p = eval₂Hom (Polynomial.C.comp (C : R →+* MvPolynomial (Fin n) R)) @@ -356,10 +355,10 @@ theorem finSuccEquiv_comp_C_eq_C {R : Type u} [CommSemiring R] (n : ℕ) : variable {n} {R} -theorem finSuccEquiv_X_zero : finSuccEquiv R n (X 0) = Polynomial.X := by simp +theorem finSuccEquiv_X_zero : finSuccEquiv R n (X 0) = Polynomial.X := by simp [finSuccEquiv_apply] theorem finSuccEquiv_X_succ {j : Fin n} : finSuccEquiv R n (X j.succ) = Polynomial.C (X j) := by - simp + simp [finSuccEquiv_apply] /-- The coefficient of `m` in the `i`-th coefficient of `finSuccEquiv R n f` equals the coefficient of `Finsupp.cons i m` in `f`. -/ @@ -414,7 +413,7 @@ theorem coeff_eval_eq_eval_coeff (s' : Fin n → R) (f : Polynomial (MvPolynomia simp only [Polynomial.coeff_map] theorem support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {m : Fin n →₀ ℕ} : - m ∈ (Polynomial.coeff ((finSuccEquiv R n) f) i).support ↔ Finsupp.cons i m ∈ f.support := by + m ∈ ((finSuccEquiv R n f).coeff i).support ↔ m.cons i ∈ f.support := by apply Iff.intro · intro h simpa [← finSuccEquiv_coeff_coeff] using h @@ -441,7 +440,7 @@ lemma totalDegree_coeff_finSuccEquiv_add_le (f : MvPolynomial (Fin (n + 1)) R) ( · rw [← support_coeff_finSuccEquiv] exact hσ1 -theorem finSuccEquiv_support (f : MvPolynomial (Fin (n + 1)) R) : +theorem support_finSuccEquiv (f : MvPolynomial (Fin (n + 1)) R) : (finSuccEquiv R n f).support = Finset.image (fun m : Fin (n + 1) →₀ ℕ => m 0) f.support := by ext i rw [Polynomial.mem_support_iff, Finset.mem_image, Finsupp.ne_iff] @@ -454,9 +453,14 @@ theorem finSuccEquiv_support (f : MvPolynomial (Fin (n + 1)) R) : refine ⟨tail m, ?_⟩ rwa [← coeff, zero_apply, ← mem_support_iff, support_coeff_finSuccEquiv, cons_tail] -theorem finSuccEquiv_support' {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : - Finset.image (Finsupp.cons i) (Polynomial.coeff ((finSuccEquiv R n) f) i).support = - f.support.filter fun m => m 0 = i := by +@[deprecated (since := "2024-11-05")] alias finSuccEquiv_support := support_finSuccEquiv + +theorem mem_support_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {x} : + x ∈ (finSuccEquiv R n f).support ↔ x ∈ (fun m : Fin (n + 1) →₀ _ ↦ m 0) '' f.support := by + simpa using congr(x ∈ $(support_finSuccEquiv f)) + +theorem image_support_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : + ((finSuccEquiv R n f).coeff i).support.image (Finsupp.cons i) = {m ∈ f.support | m 0 = i} := by ext m rw [Finset.mem_filter, Finset.mem_image, mem_support_iff] conv_lhs => @@ -472,6 +476,19 @@ theorem finSuccEquiv_support' {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} : rw [← h.2, cons_tail] simp [h.1] +@[deprecated (since := "2024-11-05")] alias finSuccEquiv_support' := image_support_finSuccEquiv + +lemma mem_image_support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {x} : + x ∈ Finsupp.cons i '' ((finSuccEquiv R n f).coeff i).support ↔ + x ∈ f.support ∧ x 0 = i := by + simpa using congr(x ∈ $image_support_finSuccEquiv) + +lemma mem_support_coeff_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} {i : ℕ} {x} : + x ∈ ((finSuccEquiv R n f).coeff i).support ↔ x.cons i ∈ f.support := by + rw [← (Finsupp.cons_right_injective i).mem_finset_image (a := x), + image_support_finSuccEquiv] + simp only [Finset.mem_filter, mem_support_iff, ne_eq, cons_zero, and_true] + -- TODO: generalize `finSuccEquiv R n` to an arbitrary ZeroHom theorem support_finSuccEquiv_nonempty {f : MvPolynomial (Fin (n + 1)) R} (h : f ≠ 0) : (finSuccEquiv R n f).support.Nonempty := by @@ -485,7 +502,7 @@ theorem degree_finSuccEquiv {f : MvPolynomial (Fin (n + 1)) R} (h : f ≠ 0) : have h₂ : WithBot.some = Nat.cast := rfl have h' : ((finSuccEquiv R n f).support.sup fun x => x) = degreeOf 0 f := by - rw [degreeOf_eq_sup, finSuccEquiv_support f, Finset.sup_image, h₀] + rw [degreeOf_eq_sup, support_finSuccEquiv, Finset.sup_image, h₀] rw [Polynomial.degree, ← h', ← h₂, Finset.coe_sup_of_nonempty (support_finSuccEquiv_nonempty h), Finset.max_eq_sup_coe, h₁] @@ -524,8 +541,8 @@ lemma finSuccEquiv_rename_finSuccEquiv (e : σ ≃ Fin n) (φ : MvPolynomial (Op (Polynomial.mapRingHom (rename e).toRingHom).comp (optionEquivLeft R σ) by exact DFunLike.congr_fun this φ apply ringHom_ext - · simp [Polynomial.algebraMap_apply, algebraMap_eq] - · rintro (i|i) <;> simp + · simp [Polynomial.algebraMap_apply, algebraMap_eq, finSuccEquiv_apply] + · rintro (i|i) <;> simp [finSuccEquiv_apply] end diff --git a/Mathlib/Data/Finsupp/Fin.lean b/Mathlib/Data/Finsupp/Fin.lean index b75198493ebe5..4e7868d955fcf 100644 --- a/Mathlib/Data/Finsupp/Fin.lean +++ b/Mathlib/Data/Finsupp/Fin.lean @@ -17,6 +17,7 @@ In this context, we prove some usual properties of `tail` and `cons`, analogous `Data.Fin.Tuple.Basic`. -/ +open Function noncomputable section @@ -86,4 +87,8 @@ lemma cons_support : (s.cons y).support ⊆ insert 0 (s.support.map (Fin.succEmb rintro i rfl simpa [Finsupp.mem_support_iff] using hi +lemma cons_right_injective {n : ℕ} {M : Type*} [Zero M] (y : M) : + Injective (Finsupp.cons y : (Fin n →₀ M) → Fin (n + 1) →₀ M) := + (equivFunOnFinite.symm.injective.comp ((Fin.cons_right_injective _).comp DFunLike.coe_injective)) + end Finsupp From e11f26c7e0afe0022f9f3f445530d54789d47542 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Mon, 25 Nov 2024 20:28:03 +0000 Subject: [PATCH 161/250] Update Batteries branch for testing https://github.com/leanprover-community/batteries/pull/1055 --- 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 aa466b92f0770..9bdd2429dc155 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de", + "rev": "6052abd696ded71dd4c23a6118b9d7b4a4704b30", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "main", + "inputRev": "finrange", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index 300ff2124f428..bcedff94bc262 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" @ git "main" +require "leanprover-community" / "batteries" from git "https://github.com/leanprover-community/batteries" @ "finrange" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.46" From 69eeb7456e3ae4dbacb7ea34191d938acb9275e7 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 20:58:34 +0000 Subject: [PATCH 162/250] chore(Data/Nat/Defs): rename two lemmas that were not about `Ne` (#19476) From LeanCamCombi --- Mathlib/Algebra/Group/Nat/Even.lean | 2 +- Mathlib/Algebra/Ring/Parity.lean | 2 +- Mathlib/Data/Nat/Defs.lean | 7 +++++-- 3 files changed, 7 insertions(+), 4 deletions(-) diff --git a/Mathlib/Algebra/Group/Nat/Even.lean b/Mathlib/Algebra/Group/Nat/Even.lean index 86fcfd3209c0d..3377025e9fb61 100644 --- a/Mathlib/Algebra/Group/Nat/Even.lean +++ b/Mathlib/Algebra/Group/Nat/Even.lean @@ -31,7 +31,7 @@ instance : DecidablePred (IsSquare : ℕ → Prop) := fun m ↦ decidable_of_iff' (Nat.sqrt m * Nat.sqrt m = m) <| by simp_rw [← Nat.exists_mul_self m, IsSquare, eq_comm] -lemma not_even_iff : ¬ Even n ↔ n % 2 = 1 := by rw [even_iff, mod_two_ne_zero] +lemma not_even_iff : ¬ Even n ↔ n % 2 = 1 := by rw [even_iff, mod_two_not_eq_zero] @[simp] lemma two_dvd_ne_zero : ¬2 ∣ n ↔ n % 2 = 1 := (even_iff_exists_two_nsmul _).symm.not.trans not_even_iff diff --git a/Mathlib/Algebra/Ring/Parity.lean b/Mathlib/Algebra/Ring/Parity.lean index 7b3959696657a..f74ff162454d5 100644 --- a/Mathlib/Algebra/Ring/Parity.lean +++ b/Mathlib/Algebra/Ring/Parity.lean @@ -201,7 +201,7 @@ lemma odd_iff : Odd n ↔ n % 2 = 1 := instance : DecidablePred (Odd : ℕ → Prop) := fun _ ↦ decidable_of_iff _ odd_iff.symm -lemma not_odd_iff : ¬Odd n ↔ n % 2 = 0 := by rw [odd_iff, mod_two_ne_one] +lemma not_odd_iff : ¬Odd n ↔ n % 2 = 0 := by rw [odd_iff, mod_two_not_eq_one] @[simp] lemma not_odd_iff_even : ¬Odd n ↔ Even n := by rw [not_odd_iff, even_iff] @[simp] lemma not_even_iff_odd : ¬Even n ↔ Odd n := by rw [not_even_iff, odd_iff] diff --git a/Mathlib/Data/Nat/Defs.lean b/Mathlib/Data/Nat/Defs.lean index 20c3db731c333..79583654cc19e 100644 --- a/Mathlib/Data/Nat/Defs.lean +++ b/Mathlib/Data/Nat/Defs.lean @@ -968,12 +968,15 @@ lemma set_induction {S : Set ℕ} (hb : 0 ∈ S) (h_ind : ∀ k : ℕ, k ∈ S attribute [simp] Nat.dvd_zero -@[simp] lemma mod_two_ne_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by +@[simp] lemma mod_two_not_eq_one : ¬n % 2 = 1 ↔ n % 2 = 0 := by cases mod_two_eq_zero_or_one n <;> simp [*] -@[simp] lemma mod_two_ne_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by +@[simp] lemma mod_two_not_eq_zero : ¬n % 2 = 0 ↔ n % 2 = 1 := by cases mod_two_eq_zero_or_one n <;> simp [*] +lemma mod_two_ne_one : n % 2 ≠ 1 ↔ n % 2 = 0 := mod_two_not_eq_one +lemma mod_two_ne_zero : n % 2 ≠ 0 ↔ n % 2 = 1 := mod_two_not_eq_zero + @[deprecated mod_mul_right_div_self (since := "2024-05-29")] lemma div_mod_eq_mod_mul_div (a b c : ℕ) : a / b % c = a % (b * c) / b := (mod_mul_right_div_self a b c).symm From 6e1b417164a9b2dba4c7f38925e8b4ba4c22e7f1 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Mon, 25 Nov 2024 21:33:59 +0000 Subject: [PATCH 163/250] chore(Pointwise): rename `image_inv` to `image_inv_eq_inv` (#19480) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This is to make space for `f '' s⁻¹ = (f '' s)⁻¹`. Also rename `smul_card_le` to `card_smul_le`. From LeanCamCombi --- .../Algebra/Group/Pointwise/Finset/Basic.lean | 16 ++++++++----- .../Algebra/Group/Pointwise/Set/Basic.lean | 24 ++++++++++++------- Mathlib/Algebra/Group/Pointwise/Set/Card.lean | 4 ++-- .../Algebra/Group/Pointwise/Set/Finite.lean | 2 +- .../GroupWithZero/Pointwise/Finset.lean | 4 ++-- .../Group/Pointwise/CompleteLattice.lean | 8 +++---- Mathlib/Algebra/Ring/Pointwise/Finset.lean | 4 ++-- Mathlib/Algebra/Ring/Pointwise/Set.lean | 4 ++-- Mathlib/Analysis/Convex/Hull.lean | 2 +- Mathlib/Analysis/Convex/Star.lean | 2 +- Mathlib/Analysis/Convex/Topology.lean | 2 +- Mathlib/Analysis/Normed/Group/Pointwise.lean | 4 ++-- Mathlib/Analysis/Normed/Group/Quotient.lean | 2 +- Mathlib/Data/Real/Cardinality.lean | 2 +- Mathlib/Data/Set/Pointwise/SMul.lean | 8 +++---- Mathlib/MeasureTheory/Group/Arithmetic.lean | 2 +- Mathlib/MeasureTheory/Group/Prod.lean | 6 ++--- .../MeasureTheory/Measure/Haar/OfBasis.lean | 2 +- Mathlib/Topology/Algebra/Field.lean | 2 +- Mathlib/Topology/Algebra/Group/Basic.lean | 2 +- 20 files changed, 56 insertions(+), 46 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 4330e3cf9abae..ac78f4f5b38b5 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -198,9 +198,7 @@ scoped[Pointwise] attribute [instance] Finset.inv Finset.neg theorem inv_def : s⁻¹ = s.image fun x => x⁻¹ := rfl -@[to_additive] -theorem image_inv : (s.image fun x => x⁻¹) = s⁻¹ := - rfl +@[to_additive] lemma image_inv_eq_inv (s : Finset α) : s.image (·⁻¹) = s⁻¹ := rfl @[to_additive] theorem mem_inv {x : α} : x ∈ s⁻¹ ↔ ∃ y ∈ s, y⁻¹ = x := @@ -279,7 +277,7 @@ variable [DecidableEq α] [InvolutiveInv α] {s : Finset α} {a : α} lemma mem_inv' : a ∈ s⁻¹ ↔ a⁻¹ ∈ s := by simp [mem_inv, inv_eq_iff_eq_inv] @[to_additive (attr := simp, norm_cast)] -theorem coe_inv (s : Finset α) : ↑s⁻¹ = (s : Set α)⁻¹ := coe_image.trans Set.image_inv +theorem coe_inv (s : Finset α) : ↑s⁻¹ = (s : Set α)⁻¹ := coe_image.trans Set.image_inv_eq_inv @[to_additive (attr := simp)] theorem card_inv (s : Finset α) : s⁻¹.card = s.card := card_image_of_injective _ inv_injective @@ -322,7 +320,10 @@ lemma coe_smul (s : Finset α) (t : Finset β) : ↑(s • t) = (s : Set α) • @[to_additive] lemma smul_mem_smul : a ∈ s → b ∈ t → a • b ∈ s • t := mem_image₂_of_mem -@[to_additive] lemma smul_card_le : (s • t).card ≤ s.card • t.card := card_image₂_le .. +@[to_additive] lemma card_smul_le : #(s • t) ≤ #s * #t := card_image₂_le .. + +@[deprecated (since := "2024-11-19")] alias smul_card_le := card_smul_le +@[deprecated (since := "2024-11-19")] alias vadd_card_le := card_vadd_le @[to_additive (attr := simp)] lemma empty_smul (t : Finset β) : (∅ : Finset α) • t = ∅ := image₂_empty_left @@ -1173,6 +1174,9 @@ theorem image_mul_left' : theorem image_mul_right' : image (· * b⁻¹) t = preimage t (· * b) (mul_left_injective _).injOn := by simp +@[to_additive] +lemma image_inv (f : F) (s : Finset α) : s⁻¹.image f = (s.image f)⁻¹ := image_comm (map_inv _) + theorem image_div : (s / t).image (f : α → β) = s.image f / t.image f := image_image₂_distrib <| map_div f @@ -1793,4 +1797,4 @@ instance Nat.decidablePred_mem_vadd_set {s : Set ℕ} [DecidablePred (· ∈ s)] fun n ↦ decidable_of_iff' (a ≤ n ∧ n - a ∈ s) <| by simp only [Set.mem_vadd_set, vadd_eq_add]; aesop -set_option linter.style.longFile 1800 +set_option linter.style.longFile 2000 diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 187539c08b148..417727b33b985 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -141,10 +141,10 @@ end One section Inv /-- The pointwise inversion of set `s⁻¹` is defined as `{x | x⁻¹ ∈ s}` in locale `Pointwise`. It is -equal to `{x⁻¹ | x ∈ s}`, see `Set.image_inv`. -/ +equal to `{x⁻¹ | x ∈ s}`, see `Set.image_inv_eq_inv`. -/ @[to_additive "The pointwise negation of set `-s` is defined as `{x | -x ∈ s}` in locale `Pointwise`. - It is equal to `{-x | x ∈ s}`, see `Set.image_neg`."] + It is equal to `{-x | x ∈ s}`, see `Set.image_neg_eq_neg`."] protected def inv [Inv α] : Inv (Set α) := ⟨preimage Inv.inv⟩ @@ -218,12 +218,12 @@ theorem Nonempty.inv (h : s.Nonempty) : s⁻¹.Nonempty := nonempty_inv.2 h @[to_additive (attr := simp)] -theorem image_inv : Inv.inv '' s = s⁻¹ := +theorem image_inv_eq_inv : (·⁻¹) '' s = s⁻¹ := congr_fun (image_eq_preimage_of_inverse inv_involutive.leftInverse inv_involutive.rightInverse) _ @[to_additive (attr := simp)] theorem inv_eq_empty : s⁻¹ = ∅ ↔ s = ∅ := by - rw [← image_inv, image_eq_empty] + rw [← image_inv_eq_inv, image_eq_empty] @[to_additive (attr := simp)] noncomputable instance involutiveInv : InvolutiveInv (Set α) where @@ -238,7 +238,8 @@ theorem inv_subset_inv : s⁻¹ ⊆ t⁻¹ ↔ s ⊆ t := theorem inv_subset : s⁻¹ ⊆ t ↔ s ⊆ t⁻¹ := by rw [← inv_subset_inv, inv_inv] @[to_additive (attr := simp)] -theorem inv_singleton (a : α) : ({a} : Set α)⁻¹ = {a⁻¹} := by rw [← image_inv, image_singleton] +theorem inv_singleton (a : α) : ({a} : Set α)⁻¹ = {a⁻¹} := by + rw [← image_inv_eq_inv, image_singleton] @[to_additive (attr := simp)] theorem inv_insert (a : α) (s : Set α) : (insert a s)⁻¹ = insert a⁻¹ s⁻¹ := by @@ -246,14 +247,14 @@ theorem inv_insert (a : α) (s : Set α) : (insert a s)⁻¹ = insert a⁻¹ s @[to_additive] theorem inv_range {ι : Sort*} {f : ι → α} : (range f)⁻¹ = range fun i => (f i)⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact (range_comp ..).symm open MulOpposite @[to_additive] theorem image_op_inv : op '' s⁻¹ = (op '' s)⁻¹ := by - simp_rw [← image_inv, Function.Semiconj.set_image op_inv s] + simp_rw [← image_inv_eq_inv, Function.Semiconj.set_image op_inv s] end InvolutiveInv @@ -1171,13 +1172,13 @@ protected theorem mul_eq_one_iff : s * t = 1 ↔ ∃ a b, s = {a} ∧ t = {b} protected noncomputable def divisionMonoid : DivisionMonoid (Set α) := { Set.monoid, Set.involutiveInv, Set.div, @Set.ZPow α _ _ _ with mul_inv_rev := fun s t => by - simp_rw [← image_inv] + simp_rw [← image_inv_eq_inv] exact image_image2_antidistrib mul_inv_rev inv_eq_of_mul := fun s t h => by obtain ⟨a, b, rfl, rfl, hab⟩ := Set.mul_eq_one_iff.1 h rw [inv_singleton, inv_eq_of_mul_eq_one_right hab] div_eq_mul_inv := fun s t => by - rw [← image_id (s / t), ← image_inv] + rw [← image_id (s / t), ← image_inv_eq_inv] exact image_image2_distrib_right div_eq_mul_inv } scoped[Pointwise] attribute [instance] Set.divisionMonoid Set.subtractionMonoid @@ -1297,6 +1298,11 @@ theorem univ_mul (ht : t.Nonempty) : (univ : Set α) * t = univ := let ⟨a, ha⟩ := ht eq_univ_of_forall fun b => ⟨b * a⁻¹, trivial, a, ha, inv_mul_cancel_right ..⟩ +@[to_additive] +lemma image_inv [DivisionMonoid β] [FunLike F α β] [MonoidHomClass F α β] (f : F) (s : Set α) : + f '' s⁻¹ = (f '' s)⁻¹ := by + rw [← image_inv_eq_inv, ← image_inv_eq_inv]; exact image_comm (map_inv _) + end Group section Mul diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean index d31f9773753a2..812782481817b 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Card.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Card.lean @@ -48,11 +48,11 @@ variable [InvolutiveInv G] @[to_additive (attr := simp)] lemma _root_.Cardinal.mk_inv (s : Set G) : #↥(s⁻¹) = #s := by - rw [← image_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] + rw [← image_inv_eq_inv, Cardinal.mk_image_eq_of_injOn _ _ inv_injective.injOn] @[to_additive (attr := simp)] lemma natCard_inv (s : Set G) : Nat.card ↥(s⁻¹) = Nat.card s := by - rw [← image_inv, Nat.card_image_of_injective inv_injective] + rw [← image_inv_eq_inv, Nat.card_image_of_injective inv_injective] @[to_additive] alias card_inv := natCard_inv diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean index 3b1d30164c535..77adb76f64340 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean @@ -112,7 +112,7 @@ section InvolutiveInv variable [InvolutiveInv α] {s : Set α} @[to_additive (attr := simp)] lemma finite_inv : s⁻¹.Finite ↔ s.Finite := by - rw [← image_inv, finite_image_iff inv_injective.injOn] + rw [← image_inv_eq_inv, finite_image_iff inv_injective.injOn] @[to_additive (attr := simp)] lemma infinite_inv : s⁻¹.Infinite ↔ s.Infinite := finite_inv.not diff --git a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean index a49b451142479..c112e0a96c541 100644 --- a/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean +++ b/Mathlib/Algebra/GroupWithZero/Pointwise/Finset.lean @@ -199,11 +199,11 @@ variable [Monoid α] [AddGroup β] [DistribMulAction α β] @[simp] lemma smul_finset_neg (a : α) (t : Finset β) : a • -t = -(a • t) := by - simp only [← image_smul, ← image_neg, Function.comp_def, image_image, smul_neg] + simp only [← image_smul, ← image_neg_eq_neg, Function.comp_def, image_image, smul_neg] @[simp] protected lemma smul_neg (s : Finset α) (t : Finset β) : s • -t = -(s • t) := by - simp_rw [← image_neg]; exact image_image₂_right_comm smul_neg + simp_rw [← image_neg_eq_neg]; exact image_image₂_right_comm smul_neg end Monoid end Finset diff --git a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean index 4a98298c00046..bdd12efab2f9b 100644 --- a/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean +++ b/Mathlib/Algebra/Order/Group/Pointwise/CompleteLattice.lean @@ -39,12 +39,12 @@ variable [Group M] [MulLeftMono M] [MulRightMono M] @[to_additive] lemma csSup_inv (hs₀ : s.Nonempty) (hs₁ : BddBelow s) : sSup s⁻¹ = (sInf s)⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact ((OrderIso.inv _).map_csInf' hs₀ hs₁).symm @[to_additive] lemma csInf_inv (hs₀ : s.Nonempty) (hs₁ : BddAbove s) : sInf s⁻¹ = (sSup s)⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact ((OrderIso.inv _).map_csSup' hs₀ hs₁).symm @[to_additive] @@ -89,12 +89,12 @@ variable [Group M] [MulLeftMono M] [MulRightMono M] @[to_additive] lemma sSup_inv (s : Set M) : sSup s⁻¹ = (sInf s)⁻¹ := by - rw [← image_inv, sSup_image] + rw [← image_inv_eq_inv, sSup_image] exact ((OrderIso.inv M).map_sInf _).symm @[to_additive] lemma sInf_inv (s : Set M) : sInf s⁻¹ = (sSup s)⁻¹ := by - rw [← image_inv, sInf_image] + rw [← image_inv_eq_inv, sInf_image] exact ((OrderIso.inv M).map_sSup _).symm @[to_additive] diff --git a/Mathlib/Algebra/Ring/Pointwise/Finset.lean b/Mathlib/Algebra/Ring/Pointwise/Finset.lean index bd2a3b4e8a7f9..f62ea470d74ea 100644 --- a/Mathlib/Algebra/Ring/Pointwise/Finset.lean +++ b/Mathlib/Algebra/Ring/Pointwise/Finset.lean @@ -58,11 +58,11 @@ variable [Ring α] [AddCommGroup β] [Module α β] [DecidableEq β] {s : Finset @[simp] lemma neg_smul_finset : -a • t = -(a • t) := by - simp only [← image_smul, ← image_neg, image_image, neg_smul, Function.comp_def] + simp only [← image_smul, ← image_neg_eq_neg, image_image, neg_smul, Function.comp_def] @[simp] protected lemma neg_smul [DecidableEq α] : -s • t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image₂_image_left_comm neg_smul end Ring diff --git a/Mathlib/Algebra/Ring/Pointwise/Set.lean b/Mathlib/Algebra/Ring/Pointwise/Set.lean index 4c844172386c8..4c75fc48d5209 100644 --- a/Mathlib/Algebra/Ring/Pointwise/Set.lean +++ b/Mathlib/Algebra/Ring/Pointwise/Set.lean @@ -29,8 +29,8 @@ namespace Set /-- `Set α` has distributive negation if `α` has. -/ protected noncomputable def hasDistribNeg [Mul α] [HasDistribNeg α] : HasDistribNeg (Set α) where __ := Set.involutiveNeg - neg_mul _ _ := by simp_rw [← image_neg]; exact image2_image_left_comm neg_mul - mul_neg _ _ := by simp_rw [← image_neg]; exact image_image2_right_comm mul_neg + neg_mul _ _ := by simp_rw [← image_neg_eq_neg]; exact image2_image_left_comm neg_mul + mul_neg _ _ := by simp_rw [← image_neg_eq_neg]; exact image_image2_right_comm mul_neg scoped[Pointwise] attribute [instance] Set.hasDistribNeg diff --git a/Mathlib/Analysis/Convex/Hull.lean b/Mathlib/Analysis/Convex/Hull.lean index 846de50849e76..d627459e1acc8 100644 --- a/Mathlib/Analysis/Convex/Hull.lean +++ b/Mathlib/Analysis/Convex/Hull.lean @@ -194,7 +194,7 @@ theorem affineSpan_convexHull (s : Set E) : affineSpan 𝕜 (convexHull 𝕜 s) exact convexHull_subset_affineSpan s theorem convexHull_neg (s : Set E) : convexHull 𝕜 (-s) = -convexHull 𝕜 s := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact AffineMap.image_convexHull (-1) _ |>.symm lemma convexHull_vadd (x : E) (s : Set E) : convexHull 𝕜 (x +ᵥ s) = x +ᵥ convexHull 𝕜 s := diff --git a/Mathlib/Analysis/Convex/Star.lean b/Mathlib/Analysis/Convex/Star.lean index ad0ee8acf1ca5..dae3f3bd26901 100644 --- a/Mathlib/Analysis/Convex/Star.lean +++ b/Mathlib/Analysis/Convex/Star.lean @@ -327,7 +327,7 @@ theorem StarConvex.affine_image (f : E →ᵃ[𝕜] F) {s : Set E} (hs : StarCon rw [Convex.combo_affine_apply hab, hy'f] theorem StarConvex.neg (hs : StarConvex 𝕜 x s) : StarConvex 𝕜 (-x) (-s) := by - rw [← image_neg] + rw [← image_neg_eq_neg] exact hs.is_linear_image IsLinearMap.isLinearMap_neg theorem StarConvex.sub (hs : StarConvex 𝕜 x s) (ht : StarConvex 𝕜 y t) : diff --git a/Mathlib/Analysis/Convex/Topology.lean b/Mathlib/Analysis/Convex/Topology.lean index 86c3bddbec48e..ac81c76f071c7 100644 --- a/Mathlib/Analysis/Convex/Topology.lean +++ b/Mathlib/Analysis/Convex/Topology.lean @@ -360,7 +360,7 @@ theorem Convex.closure_subset_image_homothety_interior_of_one_lt {s : Set E} (hs refine ⟨homothety x t⁻¹ y, hs.openSegment_interior_closure_subset_interior hx hy ?_, (AffineEquiv.homothetyUnitsMulHom x (Units.mk0 t hne)).apply_symm_apply y⟩ - rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv, + rw [openSegment_eq_image_lineMap, ← inv_one, ← inv_Ioi₀ (zero_lt_one' ℝ), ← image_inv_eq_inv, image_image, homothety_eq_lineMap] exact mem_image_of_mem _ ht diff --git a/Mathlib/Analysis/Normed/Group/Pointwise.lean b/Mathlib/Analysis/Normed/Group/Pointwise.lean index 8d853167e1370..d45a0621cf76b 100644 --- a/Mathlib/Analysis/Normed/Group/Pointwise.lean +++ b/Mathlib/Analysis/Normed/Group/Pointwise.lean @@ -39,7 +39,7 @@ theorem Bornology.IsBounded.of_mul (hst : IsBounded (s * t)) : IsBounded s ∨ I @[to_additive] theorem Bornology.IsBounded.inv : IsBounded s → IsBounded s⁻¹ := by - simp_rw [isBounded_iff_forall_norm_le', ← image_inv, forall_mem_image, norm_inv'] + simp_rw [isBounded_iff_forall_norm_le', ← image_inv_eq_inv, forall_mem_image, norm_inv'] exact id @[to_additive] @@ -58,7 +58,7 @@ open EMetric @[to_additive (attr := simp)] theorem infEdist_inv_inv (x : E) (s : Set E) : infEdist x⁻¹ s⁻¹ = infEdist x s := by - rw [← image_inv, infEdist_image isometry_inv] + rw [← image_inv_eq_inv, infEdist_image isometry_inv] @[to_additive] theorem infEdist_inv (x : E) (s : Set E) : infEdist x⁻¹ s = infEdist x s⁻¹ := by diff --git a/Mathlib/Analysis/Normed/Group/Quotient.lean b/Mathlib/Analysis/Normed/Group/Quotient.lean index fbdf446ba6f4a..0009c2a130e4e 100644 --- a/Mathlib/Analysis/Normed/Group/Quotient.lean +++ b/Mathlib/Analysis/Normed/Group/Quotient.lean @@ -150,7 +150,7 @@ theorem quotient_norm_mk_le' (S : AddSubgroup M) (m : M) : ‖(m : M ⧸ S)‖ /-- The norm of the image under the natural morphism to the quotient. -/ theorem quotient_norm_mk_eq (S : AddSubgroup M) (m : M) : ‖mk' S m‖ = sInf ((‖m + ·‖) '' S) := by - rw [mk'_apply, norm_mk, sInf_image', ← infDist_image isometry_neg, image_neg, + rw [mk'_apply, norm_mk, sInf_image', ← infDist_image isometry_neg, image_neg_eq_neg, neg_coe_set (H := S), infDist_eq_iInf] simp only [dist_eq_norm', sub_neg_eq_add, add_comm] diff --git a/Mathlib/Data/Real/Cardinality.lean b/Mathlib/Data/Real/Cardinality.lean index b41af8e41d56f..a4bda74e874be 100644 --- a/Mathlib/Data/Real/Cardinality.lean +++ b/Mathlib/Data/Real/Cardinality.lean @@ -252,7 +252,7 @@ theorem mk_Ioo_real {a b : ℝ} (h : a < b) : #(Ioo a b) = 𝔠 := by replace h := sub_pos_of_lt h have h2 : #(Inv.inv '' Ioo 0 (b - a)) ≤ #(Ioo 0 (b - a)) := mk_image_le refine le_trans ?_ h2 - rw [image_inv, inv_Ioo_0_left h, mk_Ioi_real] + rw [image_inv_eq_inv, inv_Ioo_0_left h, mk_Ioi_real] /-- The cardinality of the interval [a, b). -/ theorem mk_Ico_real {a b : ℝ} (h : a < b) : #(Ico a b) = 𝔠 := diff --git a/Mathlib/Data/Set/Pointwise/SMul.lean b/Mathlib/Data/Set/Pointwise/SMul.lean index e7008d4c67fbf..d37ecd85db188 100644 --- a/Mathlib/Data/Set/Pointwise/SMul.lean +++ b/Mathlib/Data/Set/Pointwise/SMul.lean @@ -521,11 +521,11 @@ variable [Monoid α] [AddGroup β] [DistribMulAction α β] (a : α) (s : Set α @[simp] theorem smul_set_neg : a • -t = -(a • t) := by - simp_rw [← image_smul, ← image_neg, image_image, smul_neg] + simp_rw [← image_smul, ← image_neg_eq_neg, image_image, smul_neg] @[simp] protected theorem smul_neg : s • -t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image_image2_right_comm smul_neg end Monoid @@ -546,11 +546,11 @@ variable [Ring α] [AddCommGroup β] [Module α β] (a : α) (s : Set α) (t : S @[simp] theorem neg_smul_set : -a • t = -(a • t) := by - simp_rw [← image_smul, ← image_neg, image_image, neg_smul] + simp_rw [← image_smul, ← image_neg_eq_neg, image_image, neg_smul] @[simp] protected theorem neg_smul : -s • t = -(s • t) := by - simp_rw [← image_neg] + simp_rw [← image_neg_eq_neg] exact image2_image_left_comm neg_smul end Ring diff --git a/Mathlib/MeasureTheory/Group/Arithmetic.lean b/Mathlib/MeasureTheory/Group/Arithmetic.lean index a265438e4ebd7..eaba1a41a370d 100644 --- a/Mathlib/MeasureTheory/Group/Arithmetic.lean +++ b/Mathlib/MeasureTheory/Group/Arithmetic.lean @@ -419,7 +419,7 @@ theorem MeasurableSet.inv {s : Set G} (hs : MeasurableSet s) : MeasurableSet s @[to_additive] theorem measurableEmbedding_inv [InvolutiveInv α] [MeasurableInv α] : MeasurableEmbedding (Inv.inv (α := α)) := - ⟨inv_injective, measurable_inv, fun s hs ↦ s.image_inv ▸ hs.inv⟩ + ⟨inv_injective, measurable_inv, fun s hs ↦ s.image_inv_eq_inv ▸ hs.inv⟩ end Inv diff --git a/Mathlib/MeasureTheory/Group/Prod.lean b/Mathlib/MeasureTheory/Group/Prod.lean index 8c7b2f7d77ef1..a33b49bf00b01 100644 --- a/Mathlib/MeasureTheory/Group/Prod.lean +++ b/Mathlib/MeasureTheory/Group/Prod.lean @@ -263,9 +263,9 @@ theorem ae_measure_preimage_mul_right_lt_top (hμs : μ' s ≠ ∞) : apply ae_lt_top (measurable_measure_mul_right ν' sm) have h1 := measure_mul_lintegral_eq μ' ν' sm (A⁻¹.indicator 1) (measurable_one.indicator hA.inv) rw [lintegral_indicator hA.inv] at h1 - simp_rw [Pi.one_apply, setLIntegral_one, ← image_inv, indicator_image inv_injective, image_inv, - ← indicator_mul_right _ fun x => ν' ((fun y => y * x) ⁻¹' s), Function.comp, Pi.one_apply, - mul_one] at h1 + simp_rw [Pi.one_apply, setLIntegral_one, ← image_inv_eq_inv, indicator_image inv_injective, + image_inv_eq_inv, ← indicator_mul_right _ fun x => ν' ((· * x) ⁻¹' s), Function.comp, + Pi.one_apply, mul_one] at h1 rw [← lintegral_indicator hA, ← h1] exact ENNReal.mul_ne_top hμs h3A.ne diff --git a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean index a0db8d8e207b9..9f84eee98f624 100644 --- a/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean +++ b/Mathlib/MeasureTheory/Measure/Haar/OfBasis.lean @@ -116,7 +116,7 @@ theorem parallelepiped_orthonormalBasis_one_dim (b : OrthonormalBasis ι ℝ ℝ · right simp_rw [H, parallelepiped, Algebra.id.smul_eq_mul, A] simp only [F, Finset.univ_unique, Fin.default_eq_zero, mul_neg, mul_one, Finset.sum_neg_distrib, - Finset.sum_singleton, ← image_comp, Function.comp, image_neg, neg_Icc, neg_zero] + Finset.sum_singleton, ← image_comp, Function.comp, image_neg_eq_neg, neg_Icc, neg_zero] theorem parallelepiped_eq_sum_segment (v : ι → E) : parallelepiped v = ∑ i, segment ℝ 0 (v i) := by ext diff --git a/Mathlib/Topology/Algebra/Field.lean b/Mathlib/Topology/Algebra/Field.lean index cc34c6223561e..665d8fe1857c3 100644 --- a/Mathlib/Topology/Algebra/Field.lean +++ b/Mathlib/Topology/Algebra/Field.lean @@ -52,7 +52,7 @@ def Subfield.topologicalClosure (K : Subfield α) : Subfield α := rcases eq_or_ne x 0 with (rfl | h) · rwa [inv_zero] · -- Porting note (https://github.com/leanprover-community/mathlib4/issues/11215): TODO: Lean fails to find InvMemClass instance - rw [← @inv_coe_set α (Subfield α) _ _ SubfieldClass.toInvMemClass K, ← Set.image_inv] + rw [← @inv_coe_set α (Subfield α) _ _ SubfieldClass.toInvMemClass K, ← Set.image_inv_eq_inv] exact mem_closure_image (continuousAt_inv₀ h) hx } theorem Subfield.le_topologicalClosure (s : Subfield α) : s ≤ s.topologicalClosure := diff --git a/Mathlib/Topology/Algebra/Group/Basic.lean b/Mathlib/Topology/Algebra/Group/Basic.lean index ffb45ff8d0037..4b003dcf49f11 100644 --- a/Mathlib/Topology/Algebra/Group/Basic.lean +++ b/Mathlib/Topology/Algebra/Group/Basic.lean @@ -284,7 +284,7 @@ variable [TopologicalSpace G] [InvolutiveInv G] [ContinuousInv G] {s : Set G} @[to_additive] theorem IsCompact.inv (hs : IsCompact s) : IsCompact s⁻¹ := by - rw [← image_inv] + rw [← image_inv_eq_inv] exact hs.image continuous_inv variable (G) From c5724cacdb61ddd58e20d751f235e113a70c0a67 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 09:20:28 +1100 Subject: [PATCH 164/250] fix --- Mathlib/Algebra/Module/FreeLocus.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Module/FreeLocus.lean b/Mathlib/Algebra/Module/FreeLocus.lean index 9f42786132f02..1487895fd48a4 100644 --- a/Mathlib/Algebra/Module/FreeLocus.lean +++ b/Mathlib/Algebra/Module/FreeLocus.lean @@ -126,7 +126,7 @@ lemma freeLocus_localization (S : Submonoid R) : ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization S) Rₚ Mₚ := - ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap S M) (LocalizedModule.mkLinearMap p'.primeCompl M)).extendScalarsOfIsLocalization S (Localization S) @@ -203,7 +203,7 @@ lemma isLocallyConstant_rankAtStalk_freeLocus [Module.FinitePresentation R M] : ⟨fun r r' m ↦ show algebraMap _ Rₚ (r • r') • m = _ by simp [Algebra.smul_def, ← IsScalarTower.algebraMap_apply, mul_smul]; rfl⟩ have : IsScalarTower (Localization.Away f) Rₚ Mₚ := - ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • _ by rw [← mul_smul, ← Algebra.smul_def]⟩ + ⟨fun r r' m ↦ show _ = algebraMap _ Rₚ r • r' • m by rw [← mul_smul, ← Algebra.smul_def]⟩ let l := (IsLocalizedModule.liftOfLE _ _ hp' (LocalizedModule.mkLinearMap (.powers f) M) (LocalizedModule.mkLinearMap p.asIdeal.primeCompl M)).extendScalarsOfIsLocalization (.powers f) (Localization.Away f) From a9324ec463f70aa4ddd2d35d411fbbebb417063d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 09:43:14 +1100 Subject: [PATCH 165/250] fix --- Mathlib/LinearAlgebra/Prod.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/LinearAlgebra/Prod.lean b/Mathlib/LinearAlgebra/Prod.lean index 23732c0a35284..33f885c1bc407 100644 --- a/Mathlib/LinearAlgebra/Prod.lean +++ b/Mathlib/LinearAlgebra/Prod.lean @@ -985,7 +985,8 @@ most once. Then the image of `f` is the graph of some linear map `f' : H → I`. lemma LinearMap.exists_range_eq_graph {f : G →ₛₗ[σ] H × I} (hf₁ : Surjective (Prod.fst ∘ f)) (hf : ∀ g₁ g₂, (f g₁).1 = (f g₂).1 → (f g₁).2 = (f g₂).2) : ∃ f' : H →ₗ[S] I, LinearMap.range f = LinearMap.graph f' := by - obtain ⟨f', hf'⟩ := AddMonoidHom.exists_mrange_eq_mgraph (I := I) (f := f) hf₁ hf + obtain ⟨f', hf'⟩ := + AddMonoidHom.exists_mrange_eq_mgraph (G := G) (H := H) (I := I) (f := f) hf₁ hf simp only [SetLike.ext_iff, AddMonoidHom.mem_mrange, AddMonoidHom.coe_coe, AddMonoidHom.mem_mgraph] at hf' use From 864eca313bf26ad86f866dba332fe2ddc82fe0c4 Mon Sep 17 00:00:00 2001 From: "F. G. Dorais" Date: Mon, 25 Nov 2024 15:44:52 -0500 Subject: [PATCH 166/250] chore: List.finRange moved to batteries#1055 --- Mathlib/Data/Fin/VecNotation.lean | 1 - Mathlib/Data/List/Defs.lean | 8 -------- Mathlib/Data/List/FinRange.lean | 18 +----------------- Mathlib/ModelTheory/Encoding.lean | 6 +++--- 4 files changed, 4 insertions(+), 29 deletions(-) diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 9d31131cdaf6e..1b8ca2379a123 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Data.Fin.Tuple.Basic -import Mathlib.Data.List.Defs /-! # Matrix and vector notation diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 4b86b7218cdfb..b340e9b270a42 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -492,14 +492,6 @@ end MapAccumr set_option allowUnsafeReducibility true in attribute [semireducible] Fin.foldr.loop -/-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ --- Note that we use `ofFn (fun x => x)` instead of `ofFn id` to avoid leaving `id` in the terms --- for e.g. `Fintype (Fin n)`. -def finRange (n : Nat) : List (Fin n) := ofFn (fun x => x) - --- Verify that `finRange` is semireducible. -example : finRange 3 = [0, 1, 2] := rfl - section Deprecated @[deprecated List.mem_cons (since := "2024-08-10")] diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index bf592cfe78937..83ccab9a12979 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn import Batteries.Data.List.Perm +import Batteries.Data.List.FinRange import Mathlib.Data.List.Nodup /-! @@ -25,9 +26,6 @@ variable {α : Type u} theorem finRange_eq_pmap_range (n : ℕ) : finRange n = (range n).pmap Fin.mk (by simp) := by apply List.ext_getElem <;> simp [finRange] -@[simp] -theorem finRange_zero : finRange 0 = [] := rfl - @[simp] theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := by rw [finRange_eq_pmap_range] @@ -40,10 +38,6 @@ theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := by rw [finRange_eq_pmap_range] exact (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ -@[simp] -theorem length_finRange (n : ℕ) : (finRange n).length = n := by - simp [finRange] - @[simp] theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by rw [← length_eq_zero, length_finRange] @@ -56,11 +50,6 @@ theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := by rw [finRange_eq_pmap_range] exact (List.pairwise_le_range n).pmap (by simp) (by simp) -@[simp] -theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by - simp [finRange, getElem_range, getElem_pmap] - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem theorem get_finRange {n : ℕ} {i : ℕ} (h) : (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by @@ -88,11 +77,6 @@ theorem finRange_succ_eq_map (n : ℕ) : finRange n.succ = 0 :: (finRange n).map map_map] simp only [Function.comp_def, Fin.val_succ] -theorem finRange_succ (n : ℕ) : - finRange n.succ = (finRange n |>.map Fin.castSucc |>.concat (.last _)) := by - apply map_injective_iff.mpr Fin.val_injective - simp [range_succ, Function.comp_def] - -- Porting note: `map_nth_le` moved to `List.finRange_map_get` in Data.List.Range theorem ofFn_eq_pmap {n} {f : Fin n → α} : diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 180cb86b870ed..e3832e469136b 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -82,8 +82,8 @@ theorem listDecode_encode_list (l : List (L.Term α)) : simp only [h, length_append, length_map, length_finRange, le_add_iff_nonneg_right, _root_.zero_le, ↓reduceDIte, getElem_fin, cons.injEq, func.injEq, heq_eq_eq, true_and] refine ⟨funext (fun i => ?_), ?_⟩ - · rw [List.getElem_append_left, List.getElem_map, List.getElem_finRange] - simp only [length_map, length_finRange, i.2] + · simp only [length_map, length_finRange, is_lt, getElem_append_left, getElem_map, + getElem_finRange, cast_mk, Fin.eta] · simp only [length_map, length_finRange, drop_left'] /-- An encoding of terms as lists. -/ @@ -237,7 +237,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : get?_eq_some, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ rw [get_eq_getElem, getElem_append_left, getElem_map] - · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] + · simp only [getElem_finRange, cast_mk, Fin.eta, Function.comp_apply, Sum.getLeft?_inl] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] swap From 6c6006ab93e14acc43ef1bed51c019b68259925a Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 10:28:46 +1100 Subject: [PATCH 167/250] fix test --- MathlibTest/cc.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/MathlibTest/cc.lean b/MathlibTest/cc.lean index 4d0f3eb630b40..26e0d6ac83f4d 100644 --- a/MathlibTest/cc.lean +++ b/MathlibTest/cc.lean @@ -220,17 +220,21 @@ universe u open Mathlib -axiom app : {α : Type u} → {n m : Nat} → Vector α m → Vector α n → Vector α (m + n) +axiom app : {α : Type u} → {n m : Nat} → + Mathlib.Vector α m → Mathlib.Vector α n → Mathlib.Vector α (m + n) -example (n1 n2 n3 : Nat) (v1 w1 : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : n1 = n3 → v1 = w1 → HEq w1 w1' → v2 = w2 → HEq (app v1 v2) (app w1' w2) := by cc -example (n1 n2 n3 : Nat) (v1 w1 : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 : Mathlib.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app v1 v2) (app w1' w2) := by cc -example (n1 n2 n3 : Nat) (v1 w1 v : Vector Nat n1) (w1' : Vector Nat n3) (v2 w2 w : Vector Nat n2) : +example (n1 n2 n3 : Nat) + (v1 w1 v : Mathlib.Vector Nat n1) (w1' : Mathlib.Vector Nat n3) (v2 w2 w : Mathlib.Vector Nat n2) : HEq n1 n3 → v1 = w1 → HEq w1 w1' → HEq v2 w2 → HEq (app w1' w2) (app v w) → app v1 v2 = app v w := by cc From dbb613fb80683eeae413deddb3ad5189b7c3ba81 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 10:47:54 +1100 Subject: [PATCH 168/250] fix tests --- MathlibTest/NthRewrite.lean | 2 +- MathlibTest/fractran.lean | 54 ------------------------------------- 2 files changed, 1 insertion(+), 55 deletions(-) delete mode 100644 MathlibTest/fractran.lean diff --git a/MathlibTest/NthRewrite.lean b/MathlibTest/NthRewrite.lean index 287039007dce6..3b82ffdb5998a 100644 --- a/MathlibTest/NthRewrite.lean +++ b/MathlibTest/NthRewrite.lean @@ -16,7 +16,7 @@ example [AddZeroClass G] {a : G} : a + a = a + (a + 0) := by structure F where (a : ℕ) - (v : Vector ℕ a) + (v : Mathlib.Vector ℕ a) (p : v.val = []) example (f : F) : f.v.val = [] := by diff --git a/MathlibTest/fractran.lean b/MathlibTest/fractran.lean deleted file mode 100644 index 6df5e0aa16a87..0000000000000 --- a/MathlibTest/fractran.lean +++ /dev/null @@ -1,54 +0,0 @@ -import Mathlib.Data.PNat.Defs -import Mathlib.Tactic.NormNum - -open Nat - -def FRACTRAN := List (PNat × PNat) - -namespace FRACTRAN - -def mul? (n a b : PNat) : Option PNat := - if h : (n * a % b : Nat) = 0 then - some ⟨n * a / b, Nat.div_pos (le_of_dvd (Nat.mul_pos n.2 a.2) (dvd_of_mod_eq_zero h)) b.2⟩ - else - none - -def step (f : FRACTRAN) (n : PNat) : Option PNat := - f.findSome? fun ⟨a, b⟩ => mul? n a b - -def listStep (f : FRACTRAN) (n : List PNat) : List PNat := - match n with - | [] => [] - | n :: ns => - match step f n with - | some n' => n' :: n :: ns - | none => n :: ns - -def runSteps (f : FRACTRAN) (n : PNat) (k : Nat) : List PNat := - List.range k |>.foldl (fun r _ => listStep f r) [n] - -def runFor (f : FRACTRAN) (n : PNat) : Nat → PNat - | 0 => n - | k + 1 => match step f n with - | some n' => runFor f n' k - | none => n - --- Check that `[(3, 2)]` is the program for addition. -example : (runFor [(3, 2)] ⟨2^37 * 3^42, by norm_num⟩ 1000).val = 3^(37+42) := rfl - -def PRIMEGAME : FRACTRAN := [(17, 91), (78, 85), (19, 51), (23, 38), (29, 33), (77, 29), (95, 23), - (77, 19), (1, 17), (11, 13), (13, 11), (15, 2), (1, 7), (55, 1)] - -def twoPow? (n : Nat) : Option Nat := - if _ : n = 0 then none else - if _ : n = 1 then some 0 else - if h : n % 2 = 0 then - twoPow? (n / 2) |>.map (· + 1) - else - none - -example : twoPow? 1024 = some 10 := by native_decide - -#eval runSteps PRIMEGAME ⟨2, by norm_num⟩ 100000 |>.map PNat.val|>.filterMap twoPow? - -end FRACTRAN From b50e5085f97f3825c600f239ed4a9a308572acbb Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Tue, 26 Nov 2024 03:12:44 +0000 Subject: [PATCH 169/250] chore: update Mathlib dependencies 2024-11-26 (#19497) 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 aa466b92f0770..8152b35704a8a 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "44e2d2e643fd2618b01f9a0592d7dcbd3ffa22de", + "rev": "7488499a8aad6ffada87ab6db73673d88dc04c97", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 8471a92c0fc80a9309e82362c309441139b4eeb9 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 04:40:53 +0000 Subject: [PATCH 170/250] chore: remove isStrictTotalOrder_of_linearOrder (#19493) This has been around since Lean3 core, but it is unused in Mathlib. Propose deletion. --- Mathlib.lean | 1 + Mathlib/Deprecated/Order.lean | 17 +++++++++++++++++ Mathlib/Order/Defs.lean | 5 ----- 3 files changed, 18 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Deprecated/Order.lean diff --git a/Mathlib.lean b/Mathlib.lean index 5785a8dab0524..d0155a9daf6a8 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2903,6 +2903,7 @@ import Mathlib.Deprecated.LazyList import Mathlib.Deprecated.Logic import Mathlib.Deprecated.MinMax import Mathlib.Deprecated.NatLemmas +import Mathlib.Deprecated.Order import Mathlib.Deprecated.RelClasses import Mathlib.Deprecated.Ring import Mathlib.Deprecated.Subfield diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean new file mode 100644 index 0000000000000..8e0ef3256036e --- /dev/null +++ b/Mathlib/Deprecated/Order.lean @@ -0,0 +1,17 @@ +/- +Copyright (c) 2014 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn +-/ +import Mathlib.Order.Defs + +/-! +# Deprecated instances about order structures. +-/ + +variable {α : Type*} + +@[deprecated (since := "2024-11-26")] -- unused in Mathlib +instance isStrictTotalOrder_of_linearOrder [LinearOrder α] : IsStrictTotalOrder α (· < ·) where + irrefl := lt_irrefl + trichotomous := lt_trichotomy diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index b97cdc2c979d6..8c0f11d2df7c6 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -488,11 +488,6 @@ instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidab lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) --- TODO(Leo): decide whether we should keep this instance or not -instance isStrictTotalOrder_of_linearOrder : IsStrictTotalOrder α (· < ·) where - irrefl := lt_irrefl - trichotomous := lt_trichotomy - /-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := if h : x < y then h₁ h From a8f606300eaf74a2876dc206d2f030010a5b0463 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Tue, 26 Nov 2024 05:23:35 +0000 Subject: [PATCH 171/250] chore(AlgebraicTopology): move quasicategory instances to Quasicategory dir (#19472) We add the file `Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean` and move into it the proof that any strict Segal simplicial set is a quasicategory (from `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean`). We also add the file `Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean` and move into it the proof that the nerve of a category is a quasicategory (again from `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean`). --- Mathlib.lean | 2 + .../Quasicategory/Basic.lean | 2 +- .../Quasicategory/Nerve.lean | 31 ++++++ .../Quasicategory/StrictSegal.lean | 102 ++++++++++++++++++ .../SimplicialSet/StrictSegal.lean | 80 -------------- 5 files changed, 136 insertions(+), 81 deletions(-) create mode 100644 Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean create mode 100644 Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean diff --git a/Mathlib.lean b/Mathlib.lean index d0155a9daf6a8..963adc4c3dcd2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1035,6 +1035,8 @@ import Mathlib.AlgebraicTopology.FundamentalGroupoid.Product import Mathlib.AlgebraicTopology.FundamentalGroupoid.SimplyConnected import Mathlib.AlgebraicTopology.MooreComplex import Mathlib.AlgebraicTopology.Quasicategory.Basic +import Mathlib.AlgebraicTopology.Quasicategory.Nerve +import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean index 55b27d884364c..c322b54925ac3 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Basic.lean @@ -13,7 +13,7 @@ In this file we define quasicategories, a common model of infinity categories. We show that every Kan complex is a quasicategory. -In `Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean` +In `Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean`, we show that the nerve of a category is a quasicategory. ## TODO diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean new file mode 100644 index 0000000000000..22963dfe7b25f --- /dev/null +++ b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2024 Nick Ward. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Emily Riehl, Nick Ward +-/ +import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal + +/-! +# The nerve of a category is a quasicategory + +In `AlgebraicTopology.Quasicategory.StrictSegal`, we show that any +strict Segal simplicial set is a quasicategory. +In `AlgebraicTopology.SimplicialSet.StrictSegal`, we show that the nerve of a +category satisfies the strict Segal condition. + +In this file, we prove as a direct consequence that the nerve of a category is +a quasicategory. +-/ + +universe v u + +open CategoryTheory SSet + +namespace Nerve + +/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a +category is a `Quasicategory`. -/ +instance quasicategory {C : Type u} [Category.{v} C] : + Quasicategory (nerve C) := inferInstance + +end Nerve diff --git a/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean new file mode 100644 index 0000000000000..2cecadb7f1d91 --- /dev/null +++ b/Mathlib/AlgebraicTopology/Quasicategory/StrictSegal.lean @@ -0,0 +1,102 @@ +/- +Copyright (c) 2024 Nick Ward. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin, Emily Riehl, Nick Ward +-/ +import Mathlib.AlgebraicTopology.Quasicategory.Basic +import Mathlib.AlgebraicTopology.SimplicialSet.StrictSegal + +/-! +# Strict Segal simplicial sets are quasicategories + +In `AlgebraicTopology.SimplicialSet.StrictSegal`, we define the strict Segal +condition on a simplicial set `X`. We say that `X` is strict Segal if its +simplices are uniquely determined by their spine. + +In this file, we prove that any simplicial set satisfying the strict Segal +condition is a quasicategory. +-/ + +universe u + +open CategoryTheory +open Simplicial SimplicialObject SimplexCategory + +namespace SSet.StrictSegal + +/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ +instance quasicategory {X : SSet.{u}} [StrictSegal X] : Quasicategory X := by + apply quasicategory_of_filler X + intro n i σ₀ h₀ hₙ + use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ + intro j hj + apply spineInjective + ext k + dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + let ksucc := k.succ.castSucc + obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega + · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_lt hlt] + rfl + · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] + dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] + apply congr_arg + simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, + yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, + standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, + Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] + rw [mkOfSucc_δ_gt hgt] + rfl + · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ + have hn0 : n ≠ 0 := by + rintro rfl + obtain rfl : k = 0 := by omega + fin_cases i <;> contradiction + /- We construct the triangle in the standard simplex as a 2-simplex in + the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, + we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ + let triangle : Λ[n + 2, i] _[2] := by + cases n with + | zero => contradiction + | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) + /- The interval spanning from `k` to `k + 2` is equivalently the spine + of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ + have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = + X.spine 2 (σ₀.app _ triangle) := by + ext m + dsimp [spine_arrow, Path.interval, Path.map] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, + whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, + Nat.reduceAdd, Quiver.Hom.unop_op] + cases n with + | zero => contradiction + | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl + rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] + simp only [spineToDiagonal, diagonal, spineToSimplex_spine] + rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] + apply congr_arg + simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, + whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, + uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, + len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] + ext z + cases n with + | zero => contradiction + | succ _ => + fin_cases z <;> + · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, + Quiver.Hom.unop_op, Equiv.ulift_symm_down] + rw [mkOfSucc_δ_eq heq] + rfl + +end SSet.StrictSegal diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index cef9a156cab5a..71e41b710ae25 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -3,7 +3,6 @@ Copyright (c) 2024 Emily Riehl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Emily Riehl, Joël Riou, Johan Commelin, Nick Ward -/ -import Mathlib.AlgebraicTopology.Quasicategory.Basic import Mathlib.AlgebraicTopology.SimplicialSet.Nerve import Mathlib.AlgebraicTopology.SimplicialSet.Path import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction @@ -153,81 +152,6 @@ lemma spine_δ_arrow_eq (f : Path X (n + 1)) {i : Fin n} {j : Fin (n + 2)} rw [← FunctorToTypes.map_comp_apply, ← op_comp] rw [mkOfSucc_δ_eq h, spineToSimplex_edge] -/-- Any `StrictSegal` simplicial set is a `Quasicategory`. -/ -instance : Quasicategory X := by - apply quasicategory_of_filler X - intro n i σ₀ h₀ hₙ - use spineToSimplex <| Path.map (horn.spineId i h₀ hₙ) σ₀ - intro j hj - apply spineInjective - ext k - · dsimp only [spineEquiv, spine_arrow, Function.comp_apply, Equiv.coe_fn_mk] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] - let ksucc := k.succ.castSucc - obtain hlt | hgt | heq : ksucc < j ∨ j < ksucc ∨ j = ksucc := by omega - · rw [← spine_arrow, spine_δ_arrow_lt _ hlt] - dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] - apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, - yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, - Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] - rw [mkOfSucc_δ_lt hlt] - rfl - · rw [← spine_arrow, spine_δ_arrow_gt _ hgt] - dsimp only [Path.map, spine_arrow, Fin.coe_eq_castSucc] - apply congr_arg - simp only [horn, horn.spineId, standardSimplex, uliftFunctor, Functor.comp_obj, - yoneda_obj_obj, whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, - standardSimplex.objEquiv, Equiv.ulift, Equiv.coe_fn_symm_mk, - Quiver.Hom.unop_op, horn.face_coe, Subtype.mk.injEq] - rw [mkOfSucc_δ_gt hgt] - rfl - · /- The only inner horn of `Δ[2]` does not contain the diagonal edge. -/ - have hn0 : n ≠ 0 := by - rintro rfl - obtain rfl : k = 0 := by omega - fin_cases i <;> contradiction - /- We construct the triangle in the standard simplex as a 2-simplex in - the horn. While the triangle is not contained in the inner horn `Λ[2, 1]`, - we can inhabit `Λ[n + 2, i] _[2]` by induction on `n`. -/ - let triangle : Λ[n + 2, i] _[2] := by - cases n with - | zero => contradiction - | succ _ => exact horn.primitiveTriangle i h₀ hₙ k (by omega) - /- The interval spanning from `k` to `k + 2` is equivalently the spine - of the triangle with vertices `k`, `k + 1`, and `k + 2`. -/ - have hi : ((horn.spineId i h₀ hₙ).map σ₀).interval k 2 (by omega) = - X.spine 2 (σ₀.app _ triangle) := by - ext m - dsimp [spine_arrow, Path.interval, Path.map] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality] - apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, - whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, ne_eq, - whiskering_obj_obj_map, uliftFunctor_map, yoneda_obj_map, len_mk, - Nat.reduceAdd, Quiver.Hom.unop_op] - cases n with - | zero => contradiction - | succ _ => ext x; fin_cases x <;> fin_cases m <;> rfl - rw [← spine_arrow, spine_δ_arrow_eq _ heq, hi] - simp only [spineToDiagonal, diagonal, spineToSimplex_spine] - rw [← types_comp_apply (σ₀.app _) (X.map _), ← σ₀.naturality, types_comp_apply] - apply congr_arg - simp only [horn, standardSimplex, uliftFunctor, Functor.comp_obj, - whiskering_obj_obj_obj, yoneda_obj_obj, uliftFunctor_obj, - uliftFunctor_map, whiskering_obj_obj_map, yoneda_obj_map, horn.face_coe, - len_mk, Nat.reduceAdd, Quiver.Hom.unop_op, Subtype.mk.injEq, ULift.up_inj] - ext z - cases n with - | zero => contradiction - | succ _ => - fin_cases z <;> - · simp only [standardSimplex.objEquiv, uliftFunctor_map, yoneda_obj_map, - Quiver.Hom.unop_op, Equiv.ulift_symm_down] - rw [mkOfSucc_δ_eq heq] - rfl - end StrictSegal end SSet @@ -260,8 +184,4 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ -/-- By virtue of satisfying the `StrictSegal` condition, the nerve of a -category is a `Quasicategory`. -/ -instance : Quasicategory (nerve C) := inferInstance - end Nerve From c912bed75af20c28c1a7c67461a3bc31a0a8711a Mon Sep 17 00:00:00 2001 From: Michael Stoll Date: Tue, 26 Nov 2024 05:23:36 +0000 Subject: [PATCH 172/250] chore(NumberTheory/LSeries/PrimesInAP): fix name (Lfunction -> LFunction) (#19474) --- Mathlib/NumberTheory/LSeries/PrimesInAP.lean | 56 ++++++++++---------- 1 file changed, 28 insertions(+), 28 deletions(-) diff --git a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean index 3cd0c083dd5d6..840e9ccbfa672 100644 --- a/Mathlib/NumberTheory/LSeries/PrimesInAP.lean +++ b/Mathlib/NumberTheory/LSeries/PrimesInAP.lean @@ -27,14 +27,14 @@ The main steps of the proof are as follows. agrees (on `re s > 1`) with the corresponding linear combination of negative logarithmic derivatives of Dirichlet L-functions. See `ArithmeticFunction.vonMangoldt.LSeries_residueClass_eq`. -4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux a` that is +4. Define an auxiliary function `ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux a` that is this linear combination of negative logarithmic derivatives of L-functions minus `(q.totient)⁻¹/(s-1)`, which cancels the pole at `s = 1`. - See `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux` for the statement + See `ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux` for the statement that the auxiliary function agrees with the L-series of `ArithmeticFunction.vonMangoldt.residueClass` up to the term `(q.totient)⁻¹/(s-1)`. 5. Show that the auxiliary function is continuous on `re s ≥ 1`; - see `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux`. + see `ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux`. This relies heavily on the non-vanishing of Dirichlet L-functions on the *closed* half-plane `re s ≥ 1` (`DirichletCharacter.LFunction_ne_zero_of_one_le_re`), which in turn can only be stated since we know that the L-series of a Dirichlet character @@ -51,7 +51,7 @@ The main steps of the proof are as follows. ## Definitions * `ArithmeticFunction.vonMangoldt.residueClass a` (see above). -* `ArithmeticFunction.vonMangoldt.continuousOn_LfunctionResidueClassAux` (see above). +* `ArithmeticFunction.vonMangoldt.continuousOn_LFunctionResidueClassAux` (see above). ## Main Result @@ -298,19 +298,19 @@ open Classical in Dirichlet's Theorem. On `re s > 1`, it agrees with the L-series of the von Mangoldt function restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` of the pole at `s = 1`; -see `ArithmeticFunction.vonMangoldt.eqOn_LfunctionResidueClassAux`. -/ +see `ArithmeticFunction.vonMangoldt.eqOn_LFunctionResidueClassAux`. -/ noncomputable -abbrev LfunctionResidueClassAux (s : ℂ) : ℂ := +abbrev LFunctionResidueClassAux (s : ℂ) : ℂ := (q.totient : ℂ)⁻¹ * (-deriv (LFunctionTrivChar₁ q) s / LFunctionTrivChar₁ q s - ∑ χ ∈ ({1}ᶜ : Finset (DirichletCharacter ℂ q)), χ a⁻¹ * deriv (LFunction χ) s / LFunction χ s) /-- The auxiliary function is continuous away from the zeros of the L-functions of the Dirichlet characters mod `q` (including at `s = 1`). -/ -lemma continuousOn_LfunctionResidueClassAux' : - ContinuousOn (LfunctionResidueClassAux a) +lemma continuousOn_LFunctionResidueClassAux' : + ContinuousOn (LFunctionResidueClassAux a) {s | s = 1 ∨ ∀ χ : DirichletCharacter ℂ q, LFunction χ s ≠ 0} := by - rw [show LfunctionResidueClassAux a = fun s ↦ _ from rfl] - simp only [LfunctionResidueClassAux, sub_eq_add_neg] + rw [show LFunctionResidueClassAux a = fun s ↦ _ from rfl] + simp only [LFunctionResidueClassAux, sub_eq_add_neg] refine continuousOn_const.mul <| ContinuousOn.add ?_ ?_ · refine (continuousOn_neg_logDeriv_LFunctionTrivChar₁ q).mono fun s hs ↦ ?_ have := LFunction_ne_zero_of_one_le_re (1 : DirichletCharacter ℂ q) (s := s) @@ -328,11 +328,11 @@ lemma continuousOn_LfunctionResidueClassAux' : /-- The L-series of the von Mangoldt function restricted to the prime residue class `a` mod `q` is continuous on `re s ≥ 1` except for a simple pole at `s = 1` with residue `(q.totient)⁻¹`. -The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LfunctionResidueClassAux` +The statement as given here in terms of `ArithmeticFunction.vonMangoldt.LFunctionResidueClassAux` is equivalent. -/ -lemma continuousOn_LfunctionResidueClassAux : - ContinuousOn (LfunctionResidueClassAux a) {s | 1 ≤ s.re} := by - refine (continuousOn_LfunctionResidueClassAux' a).mono fun s hs ↦ ?_ +lemma continuousOn_LFunctionResidueClassAux : + ContinuousOn (LFunctionResidueClassAux a) {s | 1 ≤ s.re} := by + refine (continuousOn_LFunctionResidueClassAux' a).mono fun s hs ↦ ?_ rcases eq_or_ne s 1 with rfl | hs₁ · simp only [ne_eq, Set.mem_setOf_eq, true_or] · simp only [ne_eq, Set.mem_setOf_eq, hs₁, false_or] @@ -345,13 +345,13 @@ open scoped LSeries.notation /-- The auxiliary function agrees on `re s > 1` with the L-series of the von Mangoldt function restricted to the residue class `a : ZMod q` minus the principal part `(q.totient)⁻¹/(s-1)` of its pole at `s = 1`. -/ -lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : - Set.EqOn (LfunctionResidueClassAux a) +lemma eqOn_LFunctionResidueClassAux (ha : IsUnit a) : + Set.EqOn (LFunctionResidueClassAux a) (fun s ↦ L ↗(residueClass a) s - (q.totient : ℂ)⁻¹ / (s - 1)) {s | 1 < s.re} := by intro s hs replace hs := Set.mem_setOf.mp hs - simp only [LSeries_residueClass_eq ha hs, LfunctionResidueClassAux] + simp only [LSeries_residueClass_eq ha hs, LFunctionResidueClassAux] rw [neg_div, ← neg_add', mul_neg, ← neg_mul, div_eq_mul_one_div (q.totient : ℂ)⁻¹, sub_eq_add_neg, ← neg_mul, ← mul_add] congrm (_ * ?_) @@ -369,9 +369,9 @@ lemma eqOn_LfunctionResidueClassAux (ha : IsUnit a) : rw [mul_comm _ 1, mul_div_mul_right _ _ <| LFunction_ne_zero_of_one_le_re 1 (.inr hs₁) hs.le] /-- The auxiliary function takes real values for real arguments `x > 1`. -/ -lemma LfunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : - LfunctionResidueClassAux a x = (LfunctionResidueClassAux a x).re := by - rw [eqOn_LfunctionResidueClassAux ha hx] +lemma LFunctionResidueClassAux_real (ha : IsUnit a) {x : ℝ} (hx : 1 < x) : + LFunctionResidueClassAux a x = (LFunctionResidueClassAux a x).re := by + rw [eqOn_LFunctionResidueClassAux ha hx] simp only [sub_re, ofReal_sub] congr 1 · rw [LSeries, re_tsum <| LSeriesSummable_of_abscissaOfAbsConv_lt_re <| @@ -394,25 +394,25 @@ lemma LSeries_residueClass_lower_bound (ha : IsUnit a) : (q.totient : ℝ)⁻¹ / (x - 1) - C ≤ ∑' n, residueClass a n / (n : ℝ) ^ x := by have H {x : ℝ} (hx : 1 < x) : ∑' n, residueClass a n / (n : ℝ) ^ x = - (LfunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by + (LFunctionResidueClassAux a x).re + (q.totient : ℝ)⁻¹ / (x - 1) := by refine ofReal_injective ?_ simp only [ofReal_tsum, ofReal_div, ofReal_cpow (Nat.cast_nonneg _), ofReal_natCast, ofReal_add, ofReal_inv, ofReal_sub, ofReal_one] - simp_rw [← LfunctionResidueClassAux_real ha hx, - eqOn_LfunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, + simp_rw [← LFunctionResidueClassAux_real ha hx, + eqOn_LFunctionResidueClassAux ha <| Set.mem_setOf.mpr (ofReal_re x ▸ hx), sub_add_cancel, LSeries, term] refine tsum_congr fun n ↦ ?_ split_ifs with hn · simp only [hn, residueClass_apply_zero, ofReal_zero, zero_div] · rfl - have : ContinuousOn (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) (Set.Icc 1 2) := - continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LfunctionResidueClassAux a) + have : ContinuousOn (fun x : ℝ ↦ (LFunctionResidueClassAux a x).re) (Set.Icc 1 2) := + continuous_re.continuousOn.comp (t := Set.univ) (continuousOn_LFunctionResidueClassAux a) (fun ⦃x⦄ a ↦ trivial) |>.comp continuous_ofReal.continuousOn fun x hx ↦ by simpa only [Set.mem_setOf_eq, ofReal_re] using hx.1 obtain ⟨C, hC⟩ := bddBelow_def.mp <| IsCompact.bddBelow_image isCompact_Icc this - replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LfunctionResidueClassAux a x).re := - hC (LfunctionResidueClassAux a x).re <| - Set.mem_image_of_mem (fun x : ℝ ↦ (LfunctionResidueClassAux a x).re) hx + replace hC {x : ℝ} (hx : x ∈ Set.Icc 1 2) : C ≤ (LFunctionResidueClassAux a x).re := + hC (LFunctionResidueClassAux a x).re <| + Set.mem_image_of_mem (fun x : ℝ ↦ (LFunctionResidueClassAux a x).re) hx refine ⟨-C, fun {x} hx ↦ ?_⟩ rw [H hx.1, add_comm, sub_neg_eq_add, add_le_add_iff_left] exact hC <| Set.mem_Icc_of_Ioc hx From b58a55482b9cbe092750d326c261e04667ef0c04 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 05:23:37 +0000 Subject: [PATCH 173/250] chore: split Order.BoundedOrder (#19488) --- .../OrderedCancelAddCommMonoidWithBounds.lean | 2 +- Mathlib.lean | 4 +- Mathlib/Algebra/Associated/Basic.lean | 2 +- .../Algebra/Order/Monoid/Canonical/Defs.lean | 2 +- .../Order/Monoid/Unbundled/TypeTags.lean | 2 +- Mathlib/Data/List/MinMax.lean | 1 + Mathlib/Data/PSigma/Order.lean | 3 +- Mathlib/Data/Prod/Lex.lean | 3 +- Mathlib/Data/Sigma/Order.lean | 3 +- .../Basic.lean} | 229 +----------------- Mathlib/Order/BoundedOrder/Lattice.lean | 165 +++++++++++++ Mathlib/Order/BoundedOrder/Monotone.lean | 124 ++++++++++ Mathlib/Order/Bounds/Basic.lean | 1 + Mathlib/Order/Disjoint.lean | 2 +- Mathlib/Order/Hom/Bounded.lean | 2 +- Mathlib/Order/Nat.lean | 2 +- Mathlib/Order/SuccPred/Limit.lean | 2 +- Mathlib/Order/WithBot.lean | 3 +- scripts/noshake.json | 1 + 19 files changed, 315 insertions(+), 238 deletions(-) rename Mathlib/Order/{BoundedOrder.lean => BoundedOrder/Basic.lean} (69%) create mode 100644 Mathlib/Order/BoundedOrder/Lattice.lean create mode 100644 Mathlib/Order/BoundedOrder/Monotone.lean diff --git a/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean index ba098aab1da69..16f749c3c471f 100644 --- a/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean +++ b/Counterexamples/OrderedCancelAddCommMonoidWithBounds.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Martin Dvorak -/ import Mathlib.Algebra.Order.Monoid.Defs -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Do not combine OrderedCancelAddCommMonoid with BoundedOrder diff --git a/Mathlib.lean b/Mathlib.lean index 963adc4c3dcd2..fddf683280feb 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3888,7 +3888,9 @@ import Mathlib.Order.BooleanAlgebra import Mathlib.Order.BooleanGenerators import Mathlib.Order.Booleanisation import Mathlib.Order.Bounded -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.BoundedOrder.Lattice +import Mathlib.Order.BoundedOrder.Monotone import Mathlib.Order.Bounds.Basic import Mathlib.Order.Bounds.Defs import Mathlib.Order.Bounds.Image diff --git a/Mathlib/Algebra/Associated/Basic.lean b/Mathlib/Algebra/Associated/Basic.lean index 91d3f4ef3b56d..4484964fb4604 100644 --- a/Mathlib/Algebra/Associated/Basic.lean +++ b/Mathlib/Algebra/Associated/Basic.lean @@ -8,9 +8,9 @@ import Mathlib.Algebra.GroupWithZero.Divisibility import Mathlib.Algebra.GroupWithZero.Hom import Mathlib.Algebra.Group.Commute.Units import Mathlib.Algebra.Group.Units.Equiv -import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Ring.Units import Mathlib.Algebra.Prime.Lemmas +import Mathlib.Order.BoundedOrder.Basic /-! # Associated elements. diff --git a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean index b18b3c81d456e..a6bc882dcc6ef 100644 --- a/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean +++ b/Mathlib/Algebra/Order/Monoid/Canonical/Defs.lean @@ -7,7 +7,7 @@ import Mathlib.Algebra.Group.Units.Basic import Mathlib.Algebra.Order.Monoid.Defs import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE import Mathlib.Algebra.NeZero -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic /-! # Canonically ordered monoids diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean index 8a7c336d6901c..761ae23ae5a20 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/TypeTags.lean @@ -4,8 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Leonardo de Moura, Mario Carneiro, Johannes Hölzl -/ import Mathlib.Algebra.Order.Monoid.Unbundled.ExistsOfLE -import Mathlib.Order.BoundedOrder import Mathlib.Algebra.Group.TypeTags.Basic +import Mathlib.Order.BoundedOrder.Basic /-! # Ordered monoid structures on `Multiplicative α` and `Additive α`. -/ diff --git a/Mathlib/Data/List/MinMax.lean b/Mathlib/Data/List/MinMax.lean index 0237438fa2eda..2a7dbf99b60f2 100644 --- a/Mathlib/Data/List/MinMax.lean +++ b/Mathlib/Data/List/MinMax.lean @@ -6,6 +6,7 @@ Authors: Minchao Wu, Chris Hughes, Mantas Bakšys import Mathlib.Data.List.Basic import Mathlib.Order.MinMax import Mathlib.Order.WithBot +import Mathlib.Order.BoundedOrder.Lattice /-! # Minimum and maximum of lists diff --git a/Mathlib/Data/PSigma/Order.lean b/Mathlib/Data/PSigma/Order.lean index c4337051e8a85..a91ce1e8563bc 100644 --- a/Mathlib/Data/PSigma/Order.lean +++ b/Mathlib/Data/PSigma/Order.lean @@ -4,10 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Minchao Wu -/ import Mathlib.Data.Sigma.Lex -import Mathlib.Order.BoundedOrder import Mathlib.Util.Notation3 import Init.NotationExtra import Mathlib.Data.Sigma.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Lexicographic order on a sigma type diff --git a/Mathlib/Data/Prod/Lex.lean b/Mathlib/Data/Prod/Lex.lean index c248e43cdc322..2aafe976d991a 100644 --- a/Mathlib/Data/Prod/Lex.lean +++ b/Mathlib/Data/Prod/Lex.lean @@ -4,7 +4,8 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison, Minchao Wu -/ import Mathlib.Data.Prod.Basic -import Mathlib.Order.BoundedOrder +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Lexicographic order diff --git a/Mathlib/Data/Sigma/Order.lean b/Mathlib/Data/Sigma/Order.lean index 3afb990febe69..7e986738b2dc0 100644 --- a/Mathlib/Data/Sigma/Order.lean +++ b/Mathlib/Data/Sigma/Order.lean @@ -4,9 +4,10 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Data.Sigma.Lex -import Mathlib.Order.BoundedOrder import Mathlib.Util.Notation3 import Mathlib.Data.Sigma.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # Orders on a sigma type diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder/Basic.lean similarity index 69% rename from Mathlib/Order/BoundedOrder.lean rename to Mathlib/Order/BoundedOrder/Basic.lean index 87bd2c47b7d17..e767d449e62e6 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder/Basic.lean @@ -3,10 +3,11 @@ Copyright (c) 2017 Johannes Hölzl. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ -import Mathlib.Order.Lattice +import Mathlib.Order.Max import Mathlib.Order.ULift import Mathlib.Tactic.PushNeg import Mathlib.Tactic.Finiteness.Attr +import Mathlib.Util.AssertExists /-! # ⊤ and ⊥, bounded lattices and variants @@ -21,15 +22,10 @@ instances for `Prop` and `fun`. * `Order α`: Order with a top/bottom element. * `BoundedOrder α`: Order with a top and bottom element. -## Common lattices - -* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` - It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and - `DistribLattice` when `OrderBot`. -* Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. - Typical examples include `Prop` and `Det α`. -/ +assert_not_exists Monotone + open Function OrderDual universe u v @@ -149,12 +145,6 @@ theorem Ne.lt_top' (h : ⊤ ≠ a) : a < ⊤ := theorem ne_top_of_le_ne_top (hb : b ≠ ⊤) (hab : a ≤ b) : a ≠ ⊤ := (hab.trans_lt hb.lt_top).ne -theorem StrictMono.apply_eq_top_iff (hf : StrictMono f) : f a = f ⊤ ↔ a = ⊤ := - ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne h, congr_arg _⟩ - -theorem StrictAnti.apply_eq_top_iff (hf : StrictAnti f) : f a = f ⊤ ↔ a = ⊤ := - ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne' h, congr_arg _⟩ - lemma top_not_mem_iff {s : Set α} : ⊤ ∉ s ↔ ∀ x ∈ s, x < ⊤ := ⟨fun h x hx ↦ Ne.lt_top (fun hx' : x = ⊤ ↦ h (hx' ▸ hx)), fun h h₀ ↦ (h ⊤ h₀).false⟩ @@ -166,14 +156,6 @@ theorem not_isMin_top : ¬IsMin (⊤ : α) := fun h => end OrderTop -theorem StrictMono.maximal_preimage_top [LinearOrder α] [Preorder β] [OrderTop β] {f : α → β} - (H : StrictMono f) {a} (h_top : f a = ⊤) (x : α) : x ≤ a := - H.maximal_of_maximal_image - (fun p => by - rw [h_top] - exact le_top) - x - theorem OrderTop.ext_top {α} {hA : PartialOrder α} (A : OrderTop α) {hB : PartialOrder α} (B : OrderTop α) (H : ∀ x y : α, (haveI := hA; x ≤ y) ↔ x ≤ y) : (@Top.top α (@OrderTop.toTop α hA.toLE A)) = (@Top.top α (@OrderTop.toTop α hB.toLE B)) := by @@ -326,12 +308,6 @@ theorem Ne.bot_lt' (h : ⊥ ≠ a) : ⊥ < a := theorem ne_bot_of_le_ne_bot (hb : b ≠ ⊥) (hab : b ≤ a) : a ≠ ⊥ := (hb.bot_lt.trans_le hab).ne' -theorem StrictMono.apply_eq_bot_iff (hf : StrictMono f) : f a = f ⊥ ↔ a = ⊥ := - hf.dual.apply_eq_top_iff - -theorem StrictAnti.apply_eq_bot_iff (hf : StrictAnti f) : f a = f ⊥ ↔ a = ⊥ := - hf.dual.apply_eq_top_iff - lemma bot_not_mem_iff {s : Set α} : ⊥ ∉ s ↔ ∀ x ∈ s, ⊥ < x := top_not_mem_iff (α := αᵒᵈ) @@ -342,14 +318,6 @@ theorem not_isMax_bot : ¬IsMax (⊥ : α) := end OrderBot -theorem StrictMono.minimal_preimage_bot [LinearOrder α] [PartialOrder β] [OrderBot β] {f : α → β} - (H : StrictMono f) {a} (h_bot : f a = ⊥) (x : α) : a ≤ x := - H.minimal_of_minimal_image - (fun p => by - rw [h_bot] - exact bot_le) - x - theorem OrderBot.ext_bot {α} {hA : PartialOrder α} (A : OrderBot α) {hB : PartialOrder α} (B : OrderBot α) (H : ∀ x y : α, (haveI := hA; x ≤ y) ↔ x ≤ y) : (@Bot.bot α (@OrderBot.toBot α hA.toLE A)) = (@Bot.bot α (@OrderBot.toBot α hB.toLE B)) := by @@ -357,65 +325,6 @@ theorem OrderBot.ext_bot {α} {hA : PartialOrder α} (A : OrderBot α) {hB : Par apply bot_unique exact @bot_le _ _ A _ -section SemilatticeSupTop - -variable [SemilatticeSup α] [OrderTop α] - --- Porting note: Not simp because simp can prove it -theorem top_sup_eq (a : α) : ⊤ ⊔ a = ⊤ := - sup_of_le_left le_top - --- Porting note: Not simp because simp can prove it -theorem sup_top_eq (a : α) : a ⊔ ⊤ = ⊤ := - sup_of_le_right le_top - -end SemilatticeSupTop - -section SemilatticeSupBot - -variable [SemilatticeSup α] [OrderBot α] {a b : α} - --- Porting note: Not simp because simp can prove it -theorem bot_sup_eq (a : α) : ⊥ ⊔ a = a := - sup_of_le_right bot_le - --- Porting note: Not simp because simp can prove it -theorem sup_bot_eq (a : α) : a ⊔ ⊥ = a := - sup_of_le_left bot_le - -@[simp] -theorem sup_eq_bot_iff : a ⊔ b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := by rw [eq_bot_iff, sup_le_iff]; simp - -end SemilatticeSupBot - -section SemilatticeInfTop - -variable [SemilatticeInf α] [OrderTop α] {a b : α} - --- Porting note: Not simp because simp can prove it -lemma top_inf_eq (a : α) : ⊤ ⊓ a = a := inf_of_le_right le_top - --- Porting note: Not simp because simp can prove it -lemma inf_top_eq (a : α) : a ⊓ ⊤ = a := inf_of_le_left le_top - -@[simp] -theorem inf_eq_top_iff : a ⊓ b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := - @sup_eq_bot_iff αᵒᵈ _ _ _ _ - -end SemilatticeInfTop - -section SemilatticeInfBot - -variable [SemilatticeInf α] [OrderBot α] - --- Porting note: Not simp because simp can prove it -lemma bot_inf_eq (a : α) : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le - --- Porting note: Not simp because simp can prove it -lemma inf_bot_eq (a : α) : a ⊓ ⊥ = ⊥ := inf_of_le_right bot_le - -end SemilatticeInfBot - /-! ### Bounded order -/ @@ -441,101 +350,8 @@ instance BoundedOrder.instSubsingleton : Subsingleton (BoundedOrder α) where end PartialOrder -section Logic - -/-! -#### In this section we prove some properties about monotone and antitone operations on `Prop` --/ - - -section Preorder - -variable [Preorder α] - -theorem monotone_and {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : - Monotone fun x => p x ∧ q x := - fun _ _ h => And.imp (m_p h) (m_q h) - --- Note: by finish [monotone] doesn't work -theorem monotone_or {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : - Monotone fun x => p x ∨ q x := - fun _ _ h => Or.imp (m_p h) (m_q h) - -theorem monotone_le {x : α} : Monotone (x ≤ ·) := fun _ _ h' h => h.trans h' - -theorem monotone_lt {x : α} : Monotone (x < ·) := fun _ _ h' h => h.trans_le h' - -theorem antitone_le {x : α} : Antitone (· ≤ x) := fun _ _ h' h => h'.trans h - -theorem antitone_lt {x : α} : Antitone (· < x) := fun _ _ h' h => h'.trans_lt h - -theorem Monotone.forall {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : - Monotone fun y => ∀ x, P x y := - fun _ _ hy h x => hP x hy <| h x - -theorem Antitone.forall {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : - Antitone fun y => ∀ x, P x y := - fun _ _ hy h x => hP x hy (h x) - -theorem Monotone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Monotone (P x)) : - Monotone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) - -theorem Antitone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Antitone (P x)) : - Antitone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) - -theorem Monotone.exists {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : - Monotone fun y => ∃ x, P x y := - fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ - -theorem Antitone.exists {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : - Antitone fun y => ∃ x, P x y := - fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ - -theorem forall_ge_iff {P : α → Prop} {x₀ : α} (hP : Monotone P) : - (∀ x ≥ x₀, P x) ↔ P x₀ := - ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ - -theorem forall_le_iff {P : α → Prop} {x₀ : α} (hP : Antitone P) : - (∀ x ≤ x₀, P x) ↔ P x₀ := - ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ - -end Preorder - -section SemilatticeSup - -variable [SemilatticeSup α] - -theorem exists_ge_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Monotone P) : - (∃ x, x₀ ≤ x ∧ P x) ↔ ∃ x, P x := - ⟨fun h => h.imp fun _ h => h.2, fun ⟨x, hx⟩ => ⟨x ⊔ x₀, le_sup_right, hP le_sup_left hx⟩⟩ - -lemma exists_and_iff_of_monotone {P Q : α → Prop} (hP : Monotone P) (hQ : Monotone Q) : - ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := - ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊔ y, ⟨hP le_sup_left hPx, hQ le_sup_right hQx⟩⟩, - fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ - -end SemilatticeSup - -section SemilatticeInf - -variable [SemilatticeInf α] - -theorem exists_le_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Antitone P) : - (∃ x, x ≤ x₀ ∧ P x) ↔ ∃ x, P x := - exists_ge_and_iff_exists <| hP.dual_left - -lemma exists_and_iff_of_antitone {P Q : α → Prop} (hP : Antitone P) (hQ : Antitone Q) : - ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := - ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊓ y, ⟨hP inf_le_left hPx, hQ inf_le_right hQx⟩⟩, - fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ - -end SemilatticeInf - -end Logic - /-! ### Function lattices -/ - namespace Pi variable {ι : Type*} {α' : ι → Type*} @@ -743,43 +559,6 @@ instance [LE α] [BoundedOrder α] : BoundedOrder (ULift.{v} α) where end ULift -section LinearOrder - -variable [LinearOrder α] - --- `simp` can prove these, so they shouldn't be simp-lemmas. -theorem min_bot_left [OrderBot α] (a : α) : min ⊥ a = ⊥ := bot_inf_eq _ - -theorem max_top_left [OrderTop α] (a : α) : max ⊤ a = ⊤ := top_sup_eq _ - -theorem min_top_left [OrderTop α] (a : α) : min ⊤ a = a := top_inf_eq _ - -theorem max_bot_left [OrderBot α] (a : α) : max ⊥ a = a := bot_sup_eq _ - -theorem min_top_right [OrderTop α] (a : α) : min a ⊤ = a := inf_top_eq _ - -theorem max_bot_right [OrderBot α] (a : α) : max a ⊥ = a := sup_bot_eq _ - -theorem min_bot_right [OrderBot α] (a : α) : min a ⊥ = ⊥ := inf_bot_eq _ - -theorem max_top_right [OrderTop α] (a : α) : max a ⊤ = ⊤ := sup_top_eq _ - -theorem max_eq_bot [OrderBot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := - sup_eq_bot_iff - -theorem min_eq_top [OrderTop α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := - inf_eq_top_iff - -@[simp] -theorem min_eq_bot [OrderBot α] {a b : α} : min a b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := by - simp_rw [← le_bot_iff, inf_le_iff] - -@[simp] -theorem max_eq_top [OrderTop α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := - @min_eq_bot αᵒᵈ _ _ a b - -end LinearOrder - section Nontrivial variable [PartialOrder α] [BoundedOrder α] [Nontrivial α] diff --git a/Mathlib/Order/BoundedOrder/Lattice.lean b/Mathlib/Order/BoundedOrder/Lattice.lean new file mode 100644 index 0000000000000..42c1cf7c0ac57 --- /dev/null +++ b/Mathlib/Order/BoundedOrder/Lattice.lean @@ -0,0 +1,165 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl +-/ +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Lattice + +/-! +# bounded lattices + +This file defines top and bottom elements (greatest and least elements) of a type, the bounded +variants of different kinds of lattices, sets up the typeclass hierarchy between them and provides +instances for `Prop` and `fun`. + +## Common lattices + +* Distributive lattices with a bottom element. Notated by `[DistribLattice α] [OrderBot α]` + It captures the properties of `Disjoint` that are common to `GeneralizedBooleanAlgebra` and + `DistribLattice` when `OrderBot`. +* Bounded and distributive lattice. Notated by `[DistribLattice α] [BoundedOrder α]`. + Typical examples include `Prop` and `Det α`. +-/ + +open Function OrderDual + +universe u v + +variable {α : Type u} {β : Type v} + +/-! ### Top, bottom element -/ + +section SemilatticeSupTop + +variable [SemilatticeSup α] [OrderTop α] + +-- Porting note: Not simp because simp can prove it +theorem top_sup_eq (a : α) : ⊤ ⊔ a = ⊤ := + sup_of_le_left le_top + +-- Porting note: Not simp because simp can prove it +theorem sup_top_eq (a : α) : a ⊔ ⊤ = ⊤ := + sup_of_le_right le_top + +end SemilatticeSupTop + +section SemilatticeSupBot + +variable [SemilatticeSup α] [OrderBot α] {a b : α} + +-- Porting note: Not simp because simp can prove it +theorem bot_sup_eq (a : α) : ⊥ ⊔ a = a := + sup_of_le_right bot_le + +-- Porting note: Not simp because simp can prove it +theorem sup_bot_eq (a : α) : a ⊔ ⊥ = a := + sup_of_le_left bot_le + +@[simp] +theorem sup_eq_bot_iff : a ⊔ b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := by rw [eq_bot_iff, sup_le_iff]; simp + +end SemilatticeSupBot + +section SemilatticeInfTop + +variable [SemilatticeInf α] [OrderTop α] {a b : α} + +-- Porting note: Not simp because simp can prove it +lemma top_inf_eq (a : α) : ⊤ ⊓ a = a := inf_of_le_right le_top + +-- Porting note: Not simp because simp can prove it +lemma inf_top_eq (a : α) : a ⊓ ⊤ = a := inf_of_le_left le_top + +@[simp] +theorem inf_eq_top_iff : a ⊓ b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := + @sup_eq_bot_iff αᵒᵈ _ _ _ _ + +end SemilatticeInfTop + +section SemilatticeInfBot + +variable [SemilatticeInf α] [OrderBot α] + +-- Porting note: Not simp because simp can prove it +lemma bot_inf_eq (a : α) : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le + +-- Porting note: Not simp because simp can prove it +lemma inf_bot_eq (a : α) : a ⊓ ⊥ = ⊥ := inf_of_le_right bot_le + +end SemilatticeInfBot + +section Logic + +/-! +#### In this section we prove some properties about monotone and antitone operations on `Prop` +-/ + +section SemilatticeSup + +variable [SemilatticeSup α] + +theorem exists_ge_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Monotone P) : + (∃ x, x₀ ≤ x ∧ P x) ↔ ∃ x, P x := + ⟨fun h => h.imp fun _ h => h.2, fun ⟨x, hx⟩ => ⟨x ⊔ x₀, le_sup_right, hP le_sup_left hx⟩⟩ + +lemma exists_and_iff_of_monotone {P Q : α → Prop} (hP : Monotone P) (hQ : Monotone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊔ y, ⟨hP le_sup_left hPx, hQ le_sup_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + +end SemilatticeSup + +section SemilatticeInf + +variable [SemilatticeInf α] + +theorem exists_le_and_iff_exists {P : α → Prop} {x₀ : α} (hP : Antitone P) : + (∃ x, x ≤ x₀ ∧ P x) ↔ ∃ x, P x := + exists_ge_and_iff_exists <| hP.dual_left + +lemma exists_and_iff_of_antitone {P Q : α → Prop} (hP : Antitone P) (hQ : Antitone Q) : + ((∃ x, P x) ∧ ∃ x, Q x) ↔ (∃ x, P x ∧ Q x) := + ⟨fun ⟨⟨x, hPx⟩, ⟨y, hQx⟩⟩ ↦ ⟨x ⊓ y, ⟨hP inf_le_left hPx, hQ inf_le_right hQx⟩⟩, + fun ⟨x, hPx, hQx⟩ ↦ ⟨⟨x, hPx⟩, ⟨x, hQx⟩⟩⟩ + +end SemilatticeInf + +end Logic + +section LinearOrder + +variable [LinearOrder α] + +-- `simp` can prove these, so they shouldn't be simp-lemmas. +theorem min_bot_left [OrderBot α] (a : α) : min ⊥ a = ⊥ := bot_inf_eq _ + +theorem max_top_left [OrderTop α] (a : α) : max ⊤ a = ⊤ := top_sup_eq _ + +theorem min_top_left [OrderTop α] (a : α) : min ⊤ a = a := top_inf_eq _ + +theorem max_bot_left [OrderBot α] (a : α) : max ⊥ a = a := bot_sup_eq _ + +theorem min_top_right [OrderTop α] (a : α) : min a ⊤ = a := inf_top_eq _ + +theorem max_bot_right [OrderBot α] (a : α) : max a ⊥ = a := sup_bot_eq _ + +theorem min_bot_right [OrderBot α] (a : α) : min a ⊥ = ⊥ := inf_bot_eq _ + +theorem max_top_right [OrderTop α] (a : α) : max a ⊤ = ⊤ := sup_top_eq _ + +theorem max_eq_bot [OrderBot α] {a b : α} : max a b = ⊥ ↔ a = ⊥ ∧ b = ⊥ := + sup_eq_bot_iff + +theorem min_eq_top [OrderTop α] {a b : α} : min a b = ⊤ ↔ a = ⊤ ∧ b = ⊤ := + inf_eq_top_iff + +@[simp] +theorem min_eq_bot [OrderBot α] {a b : α} : min a b = ⊥ ↔ a = ⊥ ∨ b = ⊥ := by + simp_rw [← le_bot_iff, inf_le_iff] + +@[simp] +theorem max_eq_top [OrderTop α] {a b : α} : max a b = ⊤ ↔ a = ⊤ ∨ b = ⊤ := + @min_eq_bot αᵒᵈ _ _ a b + +end LinearOrder diff --git a/Mathlib/Order/BoundedOrder/Monotone.lean b/Mathlib/Order/BoundedOrder/Monotone.lean new file mode 100644 index 0000000000000..6b9409de870d5 --- /dev/null +++ b/Mathlib/Order/BoundedOrder/Monotone.lean @@ -0,0 +1,124 @@ +/- +Copyright (c) 2017 Johannes Hölzl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johannes Hölzl +-/ +import Mathlib.Order.BoundedOrder.Basic +import Mathlib.Order.Monotone.Basic + +/-! +# Monotone functions on bounded orders + +-/ + +assert_not_exists SemilatticeSup + +open Function OrderDual + +universe u v + +variable {α : Type u} {β : Type v} + +/-! ### Top, bottom element -/ + +section OrderTop + +variable [PartialOrder α] [OrderTop α] [Preorder β] {f : α → β} {a b : α} + +theorem StrictMono.apply_eq_top_iff (hf : StrictMono f) : f a = f ⊤ ↔ a = ⊤ := + ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne h, congr_arg _⟩ + +theorem StrictAnti.apply_eq_top_iff (hf : StrictAnti f) : f a = f ⊤ ↔ a = ⊤ := + ⟨fun h => not_lt_top_iff.1 fun ha => (hf ha).ne' h, congr_arg _⟩ + +end OrderTop + +theorem StrictMono.maximal_preimage_top [LinearOrder α] [Preorder β] [OrderTop β] {f : α → β} + (H : StrictMono f) {a} (h_top : f a = ⊤) (x : α) : x ≤ a := + H.maximal_of_maximal_image + (fun p => by + rw [h_top] + exact le_top) + x + +section OrderBot + +variable [PartialOrder α] [OrderBot α] [Preorder β] {f : α → β} {a b : α} + +theorem StrictMono.apply_eq_bot_iff (hf : StrictMono f) : f a = f ⊥ ↔ a = ⊥ := + hf.dual.apply_eq_top_iff + +theorem StrictAnti.apply_eq_bot_iff (hf : StrictAnti f) : f a = f ⊥ ↔ a = ⊥ := + hf.dual.apply_eq_top_iff + +end OrderBot + +theorem StrictMono.minimal_preimage_bot [LinearOrder α] [PartialOrder β] [OrderBot β] {f : α → β} + (H : StrictMono f) {a} (h_bot : f a = ⊥) (x : α) : a ≤ x := + H.minimal_of_minimal_image + (fun p => by + rw [h_bot] + exact bot_le) + x + +section Logic + +/-! +#### In this section we prove some properties about monotone and antitone operations on `Prop` +-/ + + +section Preorder + +variable [Preorder α] + +theorem monotone_and {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : + Monotone fun x => p x ∧ q x := + fun _ _ h => And.imp (m_p h) (m_q h) + +-- Note: by finish [monotone] doesn't work +theorem monotone_or {p q : α → Prop} (m_p : Monotone p) (m_q : Monotone q) : + Monotone fun x => p x ∨ q x := + fun _ _ h => Or.imp (m_p h) (m_q h) + +theorem monotone_le {x : α} : Monotone (x ≤ ·) := fun _ _ h' h => h.trans h' + +theorem monotone_lt {x : α} : Monotone (x < ·) := fun _ _ h' h => h.trans_le h' + +theorem antitone_le {x : α} : Antitone (· ≤ x) := fun _ _ h' h => h'.trans h + +theorem antitone_lt {x : α} : Antitone (· < x) := fun _ _ h' h => h'.trans_lt h + +theorem Monotone.forall {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : + Monotone fun y => ∀ x, P x y := + fun _ _ hy h x => hP x hy <| h x + +theorem Antitone.forall {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : + Antitone fun y => ∀ x, P x y := + fun _ _ hy h x => hP x hy (h x) + +theorem Monotone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Monotone (P x)) : + Monotone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) + +theorem Antitone.ball {P : β → α → Prop} {s : Set β} (hP : ∀ x ∈ s, Antitone (P x)) : + Antitone fun y => ∀ x ∈ s, P x y := fun _ _ hy h x hx => hP x hx hy (h x hx) + +theorem Monotone.exists {P : β → α → Prop} (hP : ∀ x, Monotone (P x)) : + Monotone fun y => ∃ x, P x y := + fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ + +theorem Antitone.exists {P : β → α → Prop} (hP : ∀ x, Antitone (P x)) : + Antitone fun y => ∃ x, P x y := + fun _ _ hy ⟨x, hx⟩ ↦ ⟨x, hP x hy hx⟩ + +theorem forall_ge_iff {P : α → Prop} {x₀ : α} (hP : Monotone P) : + (∀ x ≥ x₀, P x) ↔ P x₀ := + ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ + +theorem forall_le_iff {P : α → Prop} {x₀ : α} (hP : Antitone P) : + (∀ x ≤ x₀, P x) ↔ P x₀ := + ⟨fun H ↦ H x₀ le_rfl, fun H _ hx ↦ hP hx H⟩ + +end Preorder + +end Logic diff --git a/Mathlib/Order/Bounds/Basic.lean b/Mathlib/Order/Bounds/Basic.lean index 99e10ab318f7a..ad2107194dde3 100644 --- a/Mathlib/Order/Bounds/Basic.lean +++ b/Mathlib/Order/Bounds/Basic.lean @@ -5,6 +5,7 @@ Authors: Johannes Hölzl, Yury Kudryashov -/ import Mathlib.Order.Bounds.Defs import Mathlib.Order.Directed +import Mathlib.Order.BoundedOrder.Monotone import Mathlib.Order.Interval.Set.Basic /-! diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index 59bcd0815347d..f291e96d7e476 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Aesop -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Disjointness and complements diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 6466c90ddaf43..feb386498fb37 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Order.Hom.Basic -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Bounded order homomorphisms diff --git a/Mathlib/Order/Nat.lean b/Mathlib/Order/Nat.lean index 2f527f28b8a53..857c598eeebea 100644 --- a/Mathlib/Order/Nat.lean +++ b/Mathlib/Order/Nat.lean @@ -3,7 +3,7 @@ Copyright (c) 2014 Floris van Doorn (c) 2016 Microsoft Corporation. All rights r Released under Apache 2.0 license as described in the file LICENSE. Authors: Floris van Doorn, Leonardo de Moura, Jeremy Avigad, Mario Carneiro -/ -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Basic /-! # The natural numbers form a linear order diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index fc51204e9d875..8929a9eebe09a 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Violeta Hernández Palacios -/ import Mathlib.Order.SuccPred.Archimedean -import Mathlib.Order.BoundedOrder +import Mathlib.Order.BoundedOrder.Lattice /-! # Successor and predecessor limits diff --git a/Mathlib/Order/WithBot.lean b/Mathlib/Order/WithBot.lean index 5be2bf90e1157..da8552cf675af 100644 --- a/Mathlib/Order/WithBot.lean +++ b/Mathlib/Order/WithBot.lean @@ -4,12 +4,13 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Logic.Nontrivial.Basic -import Mathlib.Order.BoundedOrder import Mathlib.Order.TypeTags import Mathlib.Data.Option.NAry import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Lift import Mathlib.Data.Option.Basic +import Mathlib.Order.Lattice +import Mathlib.Order.BoundedOrder.Basic /-! # `WithBot`, `WithTop` diff --git a/scripts/noshake.json b/scripts/noshake.json index 4a06c0c7680d4..b3c884f2347c8 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -337,6 +337,7 @@ "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], + "Mathlib.Order.BoundedOrder.Basic": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.Basic": ["Batteries.Tactic.Classical", "Mathlib.Tactic.GCongr.Core"], From 5e5178bc64101546e079861dbb53e784024bac2c Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Tue, 26 Nov 2024 05:38:41 +0000 Subject: [PATCH 174/250] Trigger CI for https://github.com/leanprover/lean4/pull/6200 --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 785a40b7d3e31..d6be6f1d6a9af 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "4ab397409a4f3273085129d922f6c27ebf299621", + "rev": "c8eca4fa0551087afae892e952d266f78f74357d", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "lean-pr-testing-6200", From 5b40b8f3d99cf3f3e8ce0766ef2a45bdfcf7cf4f Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 17:12:15 +1100 Subject: [PATCH 175/250] bump batteries --- 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 9bdd2429dc155..a3160f181a6cb 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,10 +5,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "6052abd696ded71dd4c23a6118b9d7b4a4704b30", + "rev": "f6d16c2a073e0385a362ad3e5d9e7e8260e298d6", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "finrange", + "inputRev": "main", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover-community/quote4", diff --git a/lakefile.lean b/lakefile.lean index bcedff94bc262..300ff2124f428 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -7,7 +7,7 @@ open Lake DSL ## Mathlib dependencies on upstream projects -/ -require "leanprover-community" / "batteries" from git "https://github.com/leanprover-community/batteries" @ "finrange" +require "leanprover-community" / "batteries" @ git "main" require "leanprover-community" / "Qq" @ git "master" require "leanprover-community" / "aesop" @ git "master" require "leanprover-community" / "proofwidgets" @ git "v0.0.46" From c94e3c95ae70b241286428eccbef27576aef040e Mon Sep 17 00:00:00 2001 From: tydeu Date: Tue, 26 Nov 2024 01:20:31 -0500 Subject: [PATCH 176/250] chore: review adaptions for leanprover/lean4#6200 --- Mathlib/Computability/Ackermann.lean | 2 +- Mathlib/Data/ENNReal/Inv.lean | 2 +- Mathlib/Data/Nat/BitIndices.lean | 2 +- Mathlib/Data/Nat/Factorization/Basic.lean | 2 +- Mathlib/Dynamics/Newton.lean | 2 +- Mathlib/NumberTheory/FactorisationProperties.lean | 2 +- Mathlib/Topology/Algebra/PontryaginDual.lean | 2 +- 7 files changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Computability/Ackermann.lean b/Mathlib/Computability/Ackermann.lean index 84d4c197c5587..92b4e6f1c3a1a 100644 --- a/Mathlib/Computability/Ackermann.lean +++ b/Mathlib/Computability/Ackermann.lean @@ -235,7 +235,7 @@ private theorem sq_le_two_pow_add_one_minus_three (n : ℕ) : n ^ 2 ≤ 2 ^ (n + norm_num apply succ_le_of_lt rw [Nat.pow_succ, mul_comm _ 2, mul_lt_mul_left (zero_lt_two' ℕ)] - exact Nat.lt_two_pow + exact Nat.lt_two_pow_self · rw [Nat.pow_succ, Nat.pow_succ] linarith [one_le_pow k 2 zero_lt_two] diff --git a/Mathlib/Data/ENNReal/Inv.lean b/Mathlib/Data/ENNReal/Inv.lean index 36003f724d580..ce120819caeaa 100644 --- a/Mathlib/Data/ENNReal/Inv.lean +++ b/Mathlib/Data/ENNReal/Inv.lean @@ -533,7 +533,7 @@ theorem exists_inv_two_pow_lt (ha : a ≠ 0) : ∃ n : ℕ, 2⁻¹ ^ n < a := by refine ⟨n, lt_trans ?_ hn⟩ rw [← ENNReal.inv_pow, ENNReal.inv_lt_inv] norm_cast - exact n.lt_two_pow + exact n.lt_two_pow_self @[simp, norm_cast] theorem coe_zpow (hr : r ≠ 0) (n : ℤ) : (↑(r ^ n) : ℝ≥0∞) = (r : ℝ≥0∞) ^ n := by diff --git a/Mathlib/Data/Nat/BitIndices.lean b/Mathlib/Data/Nat/BitIndices.lean index 659563741e3e6..8c4e3d74b89f5 100644 --- a/Mathlib/Data/Nat/BitIndices.lean +++ b/Mathlib/Data/Nat/BitIndices.lean @@ -114,6 +114,6 @@ theorem two_pow_le_of_mem_bitIndices (ha : a ∈ n.bitIndices) : 2^a ≤ n := by exact List.single_le_sum (by simp) _ <| mem_map_of_mem _ ha theorem not_mem_bitIndices_self (n : ℕ) : n ∉ n.bitIndices := - fun h ↦ (n.lt_two_pow).not_le <| two_pow_le_of_mem_bitIndices h + fun h ↦ (n.lt_two_pow_self).not_le <| two_pow_le_of_mem_bitIndices h end Nat diff --git a/Mathlib/Data/Nat/Factorization/Basic.lean b/Mathlib/Data/Nat/Factorization/Basic.lean index 73b4652c0c4a3..15659840ed9b9 100644 --- a/Mathlib/Data/Nat/Factorization/Basic.lean +++ b/Mathlib/Data/Nat/Factorization/Basic.lean @@ -356,7 +356,7 @@ theorem dvd_iff_prime_pow_dvd_dvd (n d : ℕ) : · simp rcases eq_or_ne d 0 with (rfl | hd) · simp only [zero_dvd_iff, hn, false_iff, not_forall] - exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (n.lt_two_pow).not_le⟩ + exact ⟨2, n, prime_two, dvd_zero _, mt (le_of_dvd hn.bot_lt) (n.lt_two_pow_self).not_le⟩ refine ⟨fun h p k _ hpkd => dvd_trans hpkd h, ?_⟩ rw [← factorization_prime_le_iff_dvd hd hn] intro h p pp diff --git a/Mathlib/Dynamics/Newton.lean b/Mathlib/Dynamics/Newton.lean index 3d992313b0958..3a68a8bf036e7 100644 --- a/Mathlib/Dynamics/Newton.lean +++ b/Mathlib/Dynamics/Newton.lean @@ -111,7 +111,7 @@ theorem exists_unique_nilpotent_sub_and_aeval_eq_zero · -- Existence obtain ⟨n, hn⟩ := id h refine ⟨P.newtonMap^[n] x, isNilpotent_iterate_newtonMap_sub_of_isNilpotent h n, ?_⟩ - rw [← zero_dvd_iff, ← pow_eq_zero_of_le (n.lt_two_pow).le hn] + rw [← zero_dvd_iff, ← pow_eq_zero_of_le (n.lt_two_pow_self).le hn] exact aeval_pow_two_pow_dvd_aeval_iterate_newtonMap h h' n · -- Uniqueness have ⟨u, hu⟩ := binomExpansion (P.map (algebraMap R S)) r₁ (r₂ - r₁) diff --git a/Mathlib/NumberTheory/FactorisationProperties.lean b/Mathlib/NumberTheory/FactorisationProperties.lean index 206eadd98f223..f32fbb70c0910 100644 --- a/Mathlib/NumberTheory/FactorisationProperties.lean +++ b/Mathlib/NumberTheory/FactorisationProperties.lean @@ -185,7 +185,7 @@ theorem infinite_even_deficient : {n : ℕ | Even n ∧ n.Deficient}.Infinite := constructor · exact ⟨⟨2 ^ n, by ring⟩, prime_two.deficient_pow⟩ · calc - n ≤ 2 ^ n := Nat.le_of_lt n.lt_two_pow + n ≤ 2 ^ n := Nat.le_of_lt n.lt_two_pow_self _ < 2 ^ (n + 1) := (Nat.pow_lt_pow_iff_right (Nat.one_lt_two)).mpr (lt_add_one n) theorem infinite_odd_deficient : {n : ℕ | Odd n ∧ n.Deficient}.Infinite := by diff --git a/Mathlib/Topology/Algebra/PontryaginDual.lean b/Mathlib/Topology/Algebra/PontryaginDual.lean index d2fe0b2b85cf9..2b24a7b06909d 100644 --- a/Mathlib/Topology/Algebra/PontryaginDual.lean +++ b/Mathlib/Topology/Algebra/PontryaginDual.lean @@ -73,7 +73,7 @@ instance [LocallyCompactSpace H] : LocallyCompactSpace (PontryaginDual H) := by refine lt_of_lt_of_le ht ?_ rw [div_le_iff₀' (pow_pos two_pos _), ← div_le_iff₀ hx] refine (Nat.le_ceil (Real.pi / x)).trans ?_ - exact_mod_cast (Nat.le_succ _).trans Nat.lt_two_pow.le + exact_mod_cast (Nat.le_succ _).trans Nat.lt_two_pow_self.le variable {A B C G} From a2a544ab0cc7d3e1d2a07cda2e9dedef8dbe0749 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 07:18:19 +0000 Subject: [PATCH 177/250] chore: move OrderDual.instSup/Inf up to Order.Basic (#19489) --- Mathlib/Order/Basic.lean | 6 ++++++ Mathlib/Order/Lattice.lean | 6 ------ 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 945ab1df613af..7812d48426cce 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -691,6 +691,12 @@ instance (α : Type*) [LE α] : LE αᵒᵈ := instance (α : Type*) [LT α] : LT αᵒᵈ := ⟨fun x y : α ↦ y < x⟩ +instance instSup (α : Type*) [Min α] : Max αᵒᵈ := + ⟨((· ⊓ ·) : α → α → α)⟩ + +instance instInf (α : Type*) [Max α] : Min αᵒᵈ := + ⟨((· ⊔ ·) : α → α → α)⟩ + instance instPreorder (α : Type*) [Preorder α] : Preorder αᵒᵈ where le_refl := fun _ ↦ le_refl _ le_trans := fun _ _ _ hab hbc ↦ hbc.trans hab diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index f141ab6888510..acfa2bc47f35f 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -95,12 +95,6 @@ def SemilatticeSup.mk' {α : Type*} [Max α] (sup_comm : ∀ a b : α, a ⊔ b = le_sup_right a b := by dsimp; rw [sup_comm, sup_assoc, sup_idem] sup_le a b c hac hbc := by dsimp; rwa [sup_assoc, hbc] -instance OrderDual.instSup (α : Type*) [Min α] : Max αᵒᵈ := - ⟨((· ⊓ ·) : α → α → α)⟩ - -instance OrderDual.instInf (α : Type*) [Max α] : Min αᵒᵈ := - ⟨((· ⊔ ·) : α → α → α)⟩ - section SemilatticeSup variable [SemilatticeSup α] {a b c d : α} From c8075aed3317e60c0de66103f015d93e0fd26b96 Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Tue, 26 Nov 2024 08:36:21 +0000 Subject: [PATCH 178/250] chore: bump to nightly-2024-11-26 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 1e6741ae1e185..ca6a40f05dece 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-25 +leanprover/lean4:nightly-2024-11-26 From 1ebeab334af93de65153b3fbb8708777bab5149b Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 08:42:19 +0000 Subject: [PATCH 179/250] chore: move le_refl's simp attribute to definition (#19490) --- Mathlib/Order/Basic.lean | 2 -- Mathlib/Order/Defs.lean | 2 +- 2 files changed, 1 insertion(+), 3 deletions(-) diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 7812d48426cce..84bcd9f95e9fe 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -98,8 +98,6 @@ theorem Ne.lt_of_le' : b ≠ a → a ≤ b → a < b := end PartialOrder -attribute [simp] le_refl - attribute [ext] LE alias LE.le.trans := le_trans diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean index 8c0f11d2df7c6..31c1c7508d026 100644 --- a/Mathlib/Order/Defs.lean +++ b/Mathlib/Order/Defs.lean @@ -264,7 +264,7 @@ class Preorder (α : Type*) extends LE α, LT α where variable [Preorder α] {a b c : α} /-- The relation `≤` on a preorder is reflexive. -/ -@[refl] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl +@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl /-- A version of `le_refl` where the argument is implicit -/ lemma le_rfl : a ≤ a := le_refl a From 83238e24b9113e84b6e4a66801a7dc8c4e9c8320 Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 26 Nov 2024 08:42:21 +0000 Subject: [PATCH 180/250] fix: `bors x` comments and zulip reactions (#19491) Restores the bot reacting to `bors d`, `bors merge` and `bors r+`. It only searches in `PR reviews`, but at least it should no longer fail CI. I plan to investigate making the bot search all of Zulip in a later PR. --- scripts/zulip_emoji_merge_delegate.py | 96 +++++++++++++-------------- 1 file changed, 46 insertions(+), 50 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index b279fbc6e1e11..e572aeb90e8d3 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -25,74 +25,70 @@ # Fetch the last 200 messages response = client.get_messages({ - "operator": "search", - "operand": f"https://github\.com/leanprover-community/mathlib4/pull/{PR_NUMBER}", + "anchor": "newest", + "num_before": 200, + "num_after": 0, + "narrow": [{"operator": "channel", "operand": "PR reviews"}], }) messages = response['messages'] +pr_pattern = re.compile(f'https://github.com/leanprover-community/mathlib4/pull/{PR_NUMBER}') + +print(f"Searching for: '{pr_pattern}'") + for message in messages: content = message['content'] - print(f"matched: '{message}'") # Check for emoji reactions reactions = message['reactions'] has_peace_sign = any(reaction['emoji_name'] == 'peace_sign' for reaction in reactions) has_bors = any(reaction['emoji_name'] == 'bors' for reaction in reactions) has_merge = any(reaction['emoji_name'] == 'merge' for reaction in reactions) + match = pr_pattern.search(content) + if match: + print(f"matched: '{message}'") - pr_url = f"https://api.github.com/repos/leanprover-community/mathlib4/pulls/{PR_NUMBER}" - - print('Removing peace_sign') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - print(f"result: '{result}'") - print('Removing bors') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - print(f"result: '{result}'") - - print('Removing merge') - result = client.remove_reaction({ - "message_id": message['id'], - "emoji_name": "merge" - }) - print(f"result: '{result}'") - - if 'delegated' == LABEL: - print('adding delegated') - - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "peace_sign" - }) - elif 'ready-to-merge' == LABEL: - print('adding ready-to-merge') + # removing previous emoji reactions + print("Removing previous reactions, if present.") if has_peace_sign: - client.remove_reaction({ + print('Removing peace_sign') + result = client.remove_reaction({ "message_id": message['id'], "emoji_name": "peace_sign" }) - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "bors" - }) - elif LABEL.startswith("[Merged by Bors]"): - print('adding [Merged by Bors]') - if has_peace_sign: - client.remove_reaction({ + print(f"result: '{result}'") + if has_bors: + print('Removing bors') + result = client.remove_reaction({ "message_id": message['id'], - "emoji_name": "peace_sign" + "emoji_name": "bors" }) - if has_bors: - client.remove_reaction({ + print(f"result: '{result}'") + if has_merge: + print('Removing merge') + result = client.remove_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) + print(f"result: '{result}'") + + # applying appropriate emoji reaction + print("Applying reactions, as appropriate.") + if 'ready-to-merge' == LABEL: + print('adding ready-to-merge') + client.add_reaction({ "message_id": message['id'], "emoji_name": "bors" }) - client.add_reaction({ - "message_id": message['id'], - "emoji_name": "merge" - }) + elif 'delegated' == LABEL: + print('adding delegated') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "peace_sign" + }) + elif LABEL.startswith("[Merged by Bors]"): + print('adding [Merged by Bors]') + client.add_reaction({ + "message_id": message['id'], + "emoji_name": "merge" + }) From 2973396e6752f884949148db8c063373e1d6e34e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 21:13:21 +1100 Subject: [PATCH 181/250] bump batteries --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 8b4135f271ae1..b6fd488a5ac89 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "00d092aa71e277015c11ef6bcc45d441018a642e", + "rev": "29c662a03b65054c0f4d512b81e0ac253b049a3f", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From 6c84e8f7efc21583aca881d7c9a8ce9ea25f740c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 10:41:42 +0000 Subject: [PATCH 182/250] chore: split Order.Defs (#19498) --- Mathlib.lean | 4 +- Mathlib/Algebra/Group/Commute/Defs.lean | 1 - Mathlib/Algebra/Group/Semiconj/Defs.lean | 2 +- Mathlib/Algebra/NeZero.lean | 2 +- Mathlib/Data/Bool/Basic.lean | 1 + Mathlib/Data/Char.lean | 2 +- Mathlib/Data/Int/Order/Basic.lean | 2 +- Mathlib/Data/Ordering/Lemmas.lean | 2 +- Mathlib/Deprecated/AlgebraClasses.lean | 3 +- Mathlib/Deprecated/Order.lean | 3 +- Mathlib/Logic/Basic.lean | 7 +- Mathlib/Logic/Function/Basic.lean | 2 + Mathlib/Logic/Nontrivial/Basic.lean | 2 +- Mathlib/Order/Basic.lean | 2 +- Mathlib/Order/Defs.lean | 708 ----------------------- Mathlib/Order/Defs/LinearOrder.lean | 311 ++++++++++ Mathlib/Order/Defs/PartialOrder.lean | 158 +++++ Mathlib/Order/Defs/Unbundled.lean | 278 +++++++++ Mathlib/Order/Interval/Set/Defs.lean | 2 +- Mathlib/Tactic/GCongr/Core.lean | 6 +- Mathlib/Tactic/PushNeg.lean | 2 +- MathlibTest/push_neg.lean | 2 +- scripts/noshake.json | 5 +- 23 files changed, 775 insertions(+), 732 deletions(-) delete mode 100644 Mathlib/Order/Defs.lean create mode 100644 Mathlib/Order/Defs/LinearOrder.lean create mode 100644 Mathlib/Order/Defs/PartialOrder.lean create mode 100644 Mathlib/Order/Defs/Unbundled.lean diff --git a/Mathlib.lean b/Mathlib.lean index fddf683280feb..c81570dc36eb0 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3935,7 +3935,9 @@ import Mathlib.Order.ConditionallyCompleteLattice.Indexed import Mathlib.Order.Copy import Mathlib.Order.CountableDenseLinearOrder import Mathlib.Order.Cover -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder +import Mathlib.Order.Defs.PartialOrder +import Mathlib.Order.Defs.Unbundled import Mathlib.Order.Directed import Mathlib.Order.DirectedInverseSystem import Mathlib.Order.Disjoint diff --git a/Mathlib/Algebra/Group/Commute/Defs.lean b/Mathlib/Algebra/Group/Commute/Defs.lean index 07568fa4a02eb..4d77aadfd19ab 100644 --- a/Mathlib/Algebra/Group/Commute/Defs.lean +++ b/Mathlib/Algebra/Group/Commute/Defs.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Neil Strickland, Yury Kudryashov -/ import Mathlib.Algebra.Group.Semiconj.Defs -import Mathlib.Order.Defs /-! # Commuting pairs of elements in monoids diff --git a/Mathlib/Algebra/Group/Semiconj/Defs.lean b/Mathlib/Algebra/Group/Semiconj/Defs.lean index cfeabeb0d12db..dc45b2afef400 100644 --- a/Mathlib/Algebra/Group/Semiconj/Defs.lean +++ b/Mathlib/Algebra/Group/Semiconj/Defs.lean @@ -6,7 +6,7 @@ Authors: Yury Kudryashov -- Some proofs and docs came from mathlib3 `src/algebra/commute.lean` (c) Neil Strickland import Mathlib.Algebra.Group.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled /-! # Semiconjugate elements of a semigroup diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index 3fdf3e370cb38..35ba9754734ce 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -5,7 +5,7 @@ Authors: Eric Rodriguez -/ import Mathlib.Logic.Basic import Mathlib.Algebra.Group.ZeroOne -import Mathlib.Order.Defs +import Mathlib.Order.Defs.PartialOrder /-! # `NeZero` typeclass diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 11a7e6829bea0..1b5e3ff38f835 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Jeremy Avigad -/ import Mathlib.Logic.Basic import Mathlib.Logic.Function.Defs +import Mathlib.Order.Defs.LinearOrder /-! # Booleans diff --git a/Mathlib/Data/Char.lean b/Mathlib/Data/Char.lean index 69187fc4b9eb6..ebc7f9e688601 100644 --- a/Mathlib/Data/Char.lean +++ b/Mathlib/Data/Char.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import Mathlib.Data.Nat.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder /-! # More `Char` instances diff --git a/Mathlib/Data/Int/Order/Basic.lean b/Mathlib/Data/Int/Order/Basic.lean index 4fcb9a0a8f8f8..04758c04bf6f5 100644 --- a/Mathlib/Data/Int/Order/Basic.lean +++ b/Mathlib/Data/Int/Order/Basic.lean @@ -6,7 +6,7 @@ Authors: Jeremy Avigad import Mathlib.Data.Int.Notation import Mathlib.Data.Nat.Notation -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder /-! # The order relation on the integers diff --git a/Mathlib/Data/Ordering/Lemmas.lean b/Mathlib/Data/Ordering/Lemmas.lean index 7b9714aa85ed5..42da3c30c0fb9 100644 --- a/Mathlib/Data/Ordering/Lemmas.lean +++ b/Mathlib/Data/Ordering/Lemmas.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ import Mathlib.Data.Ordering.Basic -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled /-! # Some `Ordering` lemmas diff --git a/Mathlib/Deprecated/AlgebraClasses.lean b/Mathlib/Deprecated/AlgebraClasses.lean index bb6f948871913..e9e879b6c6967 100644 --- a/Mathlib/Deprecated/AlgebraClasses.lean +++ b/Mathlib/Deprecated/AlgebraClasses.lean @@ -3,7 +3,8 @@ Copyright (c) 2017 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled +import Mathlib.Order.Defs.LinearOrder /-! # Note about deprecated files diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean index 8e0ef3256036e..9f79797c7e46f 100644 --- a/Mathlib/Deprecated/Order.lean +++ b/Mathlib/Deprecated/Order.lean @@ -3,7 +3,8 @@ Copyright (c) 2014 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.Unbundled +import Mathlib.Order.Defs.LinearOrder /-! # Deprecated instances about order structures. diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 89c3a8fe8a2f1..13c8fd29ba91a 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -6,11 +6,10 @@ Authors: Jeremy Avigad, Leonardo de Moura import Mathlib.Tactic.Attr.Register import Mathlib.Tactic.Basic import Batteries.Logic +import Batteries.Tactic.Trans import Batteries.Util.LibraryNote -import Batteries.Tactic.Lint.Basic import Mathlib.Data.Nat.Notation import Mathlib.Data.Int.Notation -import Mathlib.Order.Defs /-! # Basic logic properties @@ -133,10 +132,6 @@ section Propositional /-! ### Declarations about `implies` -/ -instance : IsRefl Prop Iff := ⟨Iff.refl⟩ - -instance : IsTrans Prop Iff := ⟨fun _ _ _ ↦ Iff.trans⟩ - alias Iff.imp := imp_congr -- This is a duplicate of `Classical.imp_iff_right_iff`. Deprecate? diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index f9b8452f77c2b..5f794b63872c6 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -8,6 +8,8 @@ import Mathlib.Logic.Basic import Mathlib.Logic.ExistsUnique import Mathlib.Logic.Nonempty import Batteries.Tactic.Init +import Mathlib.Order.Defs.PartialOrder +import Mathlib.Order.Defs.Unbundled /-! # Miscellaneous function constructions and lemmas diff --git a/Mathlib/Logic/Nontrivial/Basic.lean b/Mathlib/Logic/Nontrivial/Basic.lean index ee2fcbb4f7759..987edfff6dd39 100644 --- a/Mathlib/Logic/Nontrivial/Basic.lean +++ b/Mathlib/Logic/Nontrivial/Basic.lean @@ -7,7 +7,7 @@ import Mathlib.Data.Prod.Basic import Mathlib.Logic.Function.Basic import Mathlib.Logic.Nontrivial.Defs import Mathlib.Logic.Unique -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.Attr.Register /-! diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 84bcd9f95e9fe..caa54da5b5024 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Jeremy Avigad, Mario Carneiro -/ import Mathlib.Data.Subtype -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Order.Notation import Mathlib.Tactic.GCongr.Core import Mathlib.Tactic.Spread diff --git a/Mathlib/Order/Defs.lean b/Mathlib/Order/Defs.lean deleted file mode 100644 index 31c1c7508d026..0000000000000 --- a/Mathlib/Order/Defs.lean +++ /dev/null @@ -1,708 +0,0 @@ -/- -Copyright (c) 2016 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura --/ -import Batteries.Classes.Order -import Batteries.Tactic.Trans -import Mathlib.Data.Ordering.Basic -import Mathlib.Data.Set.Defs -import Mathlib.Tactic.ExtendDoc -import Mathlib.Tactic.Lemma -import Mathlib.Tactic.SplitIfs -import Mathlib.Tactic.TypeStar - -/-! -# Orders - -Defines classes for preorders, partial orders, and linear orders -and proves some basic lemmas about them. --/ - -/-! ### Unbundled classes -/ - -/-- An empty relation does not relate any elements. -/ -@[nolint unusedArguments] def EmptyRelation {α : Sort*} := fun _ _ : α ↦ False - -/-- `IsIrrefl X r` means the binary relation `r` on `X` is irreflexive (that is, `r x x` never -holds). -/ -class IsIrrefl (α : Sort*) (r : α → α → Prop) : Prop where - irrefl : ∀ a, ¬r a a - -/-- `IsRefl X r` means the binary relation `r` on `X` is reflexive. -/ -class IsRefl (α : Sort*) (r : α → α → Prop) : Prop where - refl : ∀ a, r a a - -/-- `IsSymm X r` means the binary relation `r` on `X` is symmetric. -/ -class IsSymm (α : Sort*) (r : α → α → Prop) : Prop where - symm : ∀ a b, r a b → r b a - -/-- `IsAsymm X r` means that the binary relation `r` on `X` is asymmetric, that is, -`r a b → ¬ r b a`. -/ -class IsAsymm (α : Sort*) (r : α → α → Prop) : Prop where - asymm : ∀ a b, r a b → ¬r b a - -/-- `IsAntisymm X r` means the binary relation `r` on `X` is antisymmetric. -/ -class IsAntisymm (α : Sort*) (r : α → α → Prop) : Prop where - antisymm : ∀ a b, r a b → r b a → a = b - -instance (priority := 100) IsAsymm.toIsAntisymm {α : Sort*} (r : α → α → Prop) [IsAsymm α r] : - IsAntisymm α r where - antisymm _ _ hx hy := (IsAsymm.asymm _ _ hx hy).elim - -/-- `IsTrans X r` means the binary relation `r` on `X` is transitive. -/ -class IsTrans (α : Sort*) (r : α → α → Prop) : Prop where - trans : ∀ a b c, r a b → r b c → r a c - -instance {α : Sort*} {r : α → α → Prop} [IsTrans α r] : Trans r r r := - ⟨IsTrans.trans _ _ _⟩ - -instance (priority := 100) {α : Sort*} {r : α → α → Prop} [Trans r r r] : IsTrans α r := - ⟨fun _ _ _ => Trans.trans⟩ - -/-- `IsTotal X r` means that the binary relation `r` on `X` is total, that is, that for any -`x y : X` we have `r x y` or `r y x`. -/ -class IsTotal (α : Sort*) (r : α → α → Prop) : Prop where - total : ∀ a b, r a b ∨ r b a - -/-- `IsPreorder X r` means that the binary relation `r` on `X` is a pre-order, that is, reflexive -and transitive. -/ -class IsPreorder (α : Sort*) (r : α → α → Prop) extends IsRefl α r, IsTrans α r : Prop - -/-- `IsPartialOrder X r` means that the binary relation `r` on `X` is a partial order, that is, -`IsPreorder X r` and `IsAntisymm X r`. -/ -class IsPartialOrder (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsAntisymm α r : Prop - -/-- `IsLinearOrder X r` means that the binary relation `r` on `X` is a linear order, that is, -`IsPartialOrder X r` and `IsTotal X r`. -/ -class IsLinearOrder (α : Sort*) (r : α → α → Prop) extends IsPartialOrder α r, IsTotal α r : Prop - -/-- `IsEquiv X r` means that the binary relation `r` on `X` is an equivalence relation, that -is, `IsPreorder X r` and `IsSymm X r`. -/ -class IsEquiv (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsSymm α r : Prop - -/-- `IsStrictOrder X r` means that the binary relation `r` on `X` is a strict order, that is, -`IsIrrefl X r` and `IsTrans X r`. -/ -class IsStrictOrder (α : Sort*) (r : α → α → Prop) extends IsIrrefl α r, IsTrans α r : Prop - -/-- `IsStrictWeakOrder X lt` means that the binary relation `lt` on `X` is a strict weak order, -that is, `IsStrictOrder X lt` and `¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a`. -/ -class IsStrictWeakOrder (α : Sort*) (lt : α → α → Prop) extends IsStrictOrder α lt : Prop where - incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a - -/-- `IsTrichotomous X lt` means that the binary relation `lt` on `X` is trichotomous, that is, -either `lt a b` or `a = b` or `lt b a` for any `a` and `b`. -/ -class IsTrichotomous (α : Sort*) (lt : α → α → Prop) : Prop where - trichotomous : ∀ a b, lt a b ∨ a = b ∨ lt b a - -/-- `IsStrictTotalOrder X lt` means that the binary relation `lt` on `X` is a strict total order, -that is, `IsTrichotomous X lt` and `IsStrictOrder X lt`. -/ -class IsStrictTotalOrder (α : Sort*) (lt : α → α → Prop) extends IsTrichotomous α lt, - IsStrictOrder α lt : Prop - -/-- Equality is an equivalence relation. -/ -instance eq_isEquiv (α : Sort*) : IsEquiv α (· = ·) where - symm := @Eq.symm _ - trans := @Eq.trans _ - refl := Eq.refl - -section - -variable {α : Sort*} {r : α → α → Prop} {a b c : α} - -/-- Local notation for an arbitrary binary relation `r`. -/ -local infixl:50 " ≺ " => r - -lemma irrefl [IsIrrefl α r] (a : α) : ¬a ≺ a := IsIrrefl.irrefl a -lemma refl [IsRefl α r] (a : α) : a ≺ a := IsRefl.refl a -lemma trans [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := IsTrans.trans _ _ _ -lemma symm [IsSymm α r] : a ≺ b → b ≺ a := IsSymm.symm _ _ -lemma antisymm [IsAntisymm α r] : a ≺ b → b ≺ a → a = b := IsAntisymm.antisymm _ _ -lemma asymm [IsAsymm α r] : a ≺ b → ¬b ≺ a := IsAsymm.asymm _ _ - -lemma trichotomous [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := - IsTrichotomous.trichotomous - -instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrrefl α r] : - IsAsymm α r := - ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ - -instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : - IsIrrefl α (fun a b => decide (r a b) = true) where - irrefl := fun a => by simpa using irrefl a - -instance IsRefl.decide [DecidableRel r] [IsRefl α r] : - IsRefl α (fun a b => decide (r a b) = true) where - refl := fun a => by simpa using refl a - -instance IsTrans.decide [DecidableRel r] [IsTrans α r] : - IsTrans α (fun a b => decide (r a b) = true) where - trans := fun a b c => by simpa using trans a b c - -instance IsSymm.decide [DecidableRel r] [IsSymm α r] : - IsSymm α (fun a b => decide (r a b) = true) where - symm := fun a b => by simpa using symm a b - -instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : - IsAntisymm α (fun a b => decide (r a b) = true) where - antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) - -instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : - IsAsymm α (fun a b => decide (r a b) = true) where - asymm := fun a b => by simpa using asymm a b - -instance IsTotal.decide [DecidableRel r] [IsTotal α r] : - IsTotal α (fun a b => decide (r a b) = true) where - total := fun a b => by simpa using total a b - -instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : - IsTrichotomous α (fun a b => decide (r a b) = true) where - trichotomous := fun a b => by simpa using trichotomous a b - -variable (r) - -@[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a -@[elab_without_expected_type] lemma refl_of [IsRefl α r] (a : α) : a ≺ a := refl a -@[elab_without_expected_type] lemma trans_of [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := _root_.trans -@[elab_without_expected_type] lemma symm_of [IsSymm α r] : a ≺ b → b ≺ a := symm -@[elab_without_expected_type] lemma asymm_of [IsAsymm α r] : a ≺ b → ¬b ≺ a := asymm - -@[elab_without_expected_type] -lemma total_of [IsTotal α r] (a b : α) : a ≺ b ∨ b ≺ a := IsTotal.total _ _ - -@[elab_without_expected_type] -lemma trichotomous_of [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := trichotomous - -section - -/-- `IsRefl` as a definition, suitable for use in proofs. -/ -def Reflexive := ∀ x, x ≺ x - -/-- `IsSymm` as a definition, suitable for use in proofs. -/ -def Symmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x - -/-- `IsTrans` as a definition, suitable for use in proofs. -/ -def Transitive := ∀ ⦃x y z⦄, x ≺ y → y ≺ z → x ≺ z - -/-- `IsIrrefl` as a definition, suitable for use in proofs. -/ -def Irreflexive := ∀ x, ¬x ≺ x - -/-- `IsAntisymm` as a definition, suitable for use in proofs. -/ -def AntiSymmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x → x = y - -/-- `IsTotal` as a definition, suitable for use in proofs. -/ -def Total := ∀ x y, x ≺ y ∨ y ≺ x - -@[deprecated Equivalence.refl (since := "2024-09-13")] -theorem Equivalence.reflexive (h : Equivalence r) : Reflexive r := h.refl - -@[deprecated Equivalence.symm (since := "2024-09-13")] -theorem Equivalence.symmetric (h : Equivalence r) : Symmetric r := - fun _ _ ↦ h.symm - -@[deprecated Equivalence.trans (since := "2024-09-13")] -theorem Equivalence.transitive (h : Equivalence r) : Transitive r := - fun _ _ _ ↦ h.trans - -variable {β : Sort*} (r : β → β → Prop) (f : α → β) - -@[deprecated "No deprecation message was provided." (since := "2024-09-13")] -theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := - fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ - -@[deprecated "No deprecation message was provided." (since := "2024-09-13")] -theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := - fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ - -end - -end - -/-! ### Minimal and maximal -/ - -section LE - -variable {α : Type*} [LE α] {P : α → Prop} {x y : α} - -/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/ -def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y - -/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/ -def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x - -lemma Minimal.prop (h : Minimal P x) : P x := - h.1 - -lemma Maximal.prop (h : Maximal P x) : P x := - h.1 - -lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y := - h.2 hy hle - -lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x := - h.2 hy hge - -end LE - -/-! ### Bundled classes -/ - -variable {α : Type*} - -section Preorder - -/-! -### Definition of `Preorder` and lemmas about types with a `Preorder` --/ - -/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/ -class Preorder (α : Type*) extends LE α, LT α where - le_refl : ∀ a : α, a ≤ a - le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c - lt := fun a b => a ≤ b ∧ ¬b ≤ a - lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl - -variable [Preorder α] {a b c : α} - -/-- The relation `≤` on a preorder is reflexive. -/ -@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl - -/-- A version of `le_refl` where the argument is implicit -/ -lemma le_rfl : a ≤ a := le_refl a - -/-- The relation `≤` on a preorder is transitive. -/ -@[trans] lemma le_trans : a ≤ b → b ≤ c → a ≤ c := Preorder.le_trans _ _ _ - -lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_not_le _ _ - -lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ - -@[deprecated "No deprecation message was provided." (since := "2024-07-30")] -theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a - | _a, _b, hab => lt_iff_le_not_le.mp hab - -lemma le_of_eq (hab : a = b) : a ≤ b := by rw [hab] -lemma le_of_lt (hab : a < b) : a ≤ b := (lt_iff_le_not_le.1 hab).1 -lemma not_le_of_lt (hab : a < b) : ¬ b ≤ a := (lt_iff_le_not_le.1 hab).2 -lemma not_le_of_gt (hab : a > b) : ¬a ≤ b := not_le_of_lt hab -lemma not_lt_of_le (hab : a ≤ b) : ¬ b < a := imp_not_comm.1 not_le_of_lt hab -lemma not_lt_of_ge (hab : a ≥ b) : ¬a < b := not_lt_of_le hab - -alias LT.lt.not_le := not_le_of_lt -alias LE.le.not_lt := not_lt_of_le - -@[trans] lemma ge_trans : a ≥ b → b ≥ c → a ≥ c := fun h₁ h₂ => le_trans h₂ h₁ - -lemma lt_irrefl (a : α) : ¬a < a := fun h ↦ not_le_of_lt h le_rfl -lemma gt_irrefl (a : α) : ¬a > a := lt_irrefl _ - -@[trans] lemma lt_of_lt_of_le (hab : a < b) (hbc : b ≤ c) : a < c := - lt_of_le_not_le (le_trans (le_of_lt hab) hbc) fun hca ↦ not_le_of_lt hab (le_trans hbc hca) - -@[trans] lemma lt_of_le_of_lt (hab : a ≤ b) (hbc : b < c) : a < c := - lt_of_le_not_le (le_trans hab (le_of_lt hbc)) fun hca ↦ not_le_of_lt hbc (le_trans hca hab) - -@[trans] lemma gt_of_gt_of_ge (h₁ : a > b) (h₂ : b ≥ c) : a > c := lt_of_le_of_lt h₂ h₁ -@[trans] lemma gt_of_ge_of_gt (h₁ : a ≥ b) (h₂ : b > c) : a > c := lt_of_lt_of_le h₂ h₁ - -@[trans] lemma lt_trans (hab : a < b) (hbc : b < c) : a < c := lt_of_lt_of_le hab (le_of_lt hbc) -@[trans] lemma gt_trans : a > b → b > c → a > c := fun h₁ h₂ => lt_trans h₂ h₁ - -lemma ne_of_lt (h : a < b) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) -lemma ne_of_gt (h : b < a) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) -lemma lt_asymm (h : a < b) : ¬b < a := fun h1 : b < a => lt_irrefl a (lt_trans h h1) - -alias not_lt_of_gt := lt_asymm -alias not_lt_of_lt := lt_asymm - -lemma le_of_lt_or_eq (h : a < b ∨ a = b) : a ≤ b := h.elim le_of_lt le_of_eq -lemma le_of_eq_or_lt (h : a = b ∨ a < b) : a ≤ b := h.elim le_of_eq le_of_lt - -instance (priority := 900) : @Trans α α α LE.le LE.le LE.le := ⟨le_trans⟩ -instance (priority := 900) : @Trans α α α LT.lt LT.lt LT.lt := ⟨lt_trans⟩ -instance (priority := 900) : @Trans α α α LT.lt LE.le LT.lt := ⟨lt_of_lt_of_le⟩ -instance (priority := 900) : @Trans α α α LE.le LT.lt LT.lt := ⟨lt_of_le_of_lt⟩ -instance (priority := 900) : @Trans α α α GE.ge GE.ge GE.ge := ⟨ge_trans⟩ -instance (priority := 900) : @Trans α α α GT.gt GT.gt GT.gt := ⟨gt_trans⟩ -instance (priority := 900) : @Trans α α α GT.gt GE.ge GT.gt := ⟨gt_of_gt_of_ge⟩ -instance (priority := 900) : @Trans α α α GE.ge GT.gt GT.gt := ⟨gt_of_ge_of_gt⟩ - -/-- `<` is decidable if `≤` is. -/ -def decidableLTOfDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel α (· < ·) - | a, b => - if hab : a ≤ b then - if hba : b ≤ a then isFalse fun hab' => not_le_of_gt hab' hba - else isTrue <| lt_of_le_not_le hab hba - else isFalse fun hab' => hab (le_of_lt hab') - -end Preorder - -section PartialOrder - -/-! -### Definition of `PartialOrder` and lemmas about types with a partial order --/ - -/-- A partial order is a reflexive, transitive, antisymmetric relation `≤`. -/ -class PartialOrder (α : Type*) extends Preorder α where - le_antisymm : ∀ a b : α, a ≤ b → b ≤ a → a = b - -variable [PartialOrder α] {a b : α} - -lemma le_antisymm : a ≤ b → b ≤ a → a = b := PartialOrder.le_antisymm _ _ - -alias eq_of_le_of_le := le_antisymm - -lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := - ⟨fun e => ⟨le_of_eq e, le_of_eq e.symm⟩, fun ⟨h1, h2⟩ => le_antisymm h1 h2⟩ - -lemma lt_of_le_of_ne : a ≤ b → a ≠ b → a < b := fun h₁ h₂ => - lt_of_le_not_le h₁ <| mt (le_antisymm h₁) h₂ - -/-- Equality is decidable if `≤` is. -/ -def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α - | a, b => - if hab : a ≤ b then - if hba : b ≤ a then isTrue (le_antisymm hab hba) else isFalse fun heq => hba (heq ▸ le_refl _) - else isFalse fun heq => hab (heq ▸ le_refl _) - -namespace Decidable - -variable [@DecidableRel α (· ≤ ·)] - -lemma lt_or_eq_of_le (hab : a ≤ b) : a < b ∨ a = b := - if hba : b ≤ a then Or.inr (le_antisymm hab hba) else Or.inl (lt_of_le_not_le hab hba) - -lemma eq_or_lt_of_le (hab : a ≤ b) : a = b ∨ a < b := - (lt_or_eq_of_le hab).symm - -lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := - ⟨lt_or_eq_of_le, le_of_lt_or_eq⟩ - -end Decidable - -attribute [local instance] Classical.propDecidable - -lemma lt_or_eq_of_le : a ≤ b → a < b ∨ a = b := Decidable.lt_or_eq_of_le -lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := Decidable.le_iff_lt_or_eq - -end PartialOrder - -section LinearOrder - -/-! -### Definition of `LinearOrder` and lemmas about types with a linear order --/ - -/-- Default definition of `max`. -/ -def maxDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := - if a ≤ b then b else a - -/-- Default definition of `min`. -/ -def minDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := - if a ≤ b then a else b - -/-- This attempts to prove that a given instance of `compare` is equal to `compareOfLessAndEq` by -introducing the arguments and trying the following approaches in order: - -1. seeing if `rfl` works -2. seeing if the `compare` at hand is nonetheless essentially `compareOfLessAndEq`, but, because of -implicit arguments, requires us to unfold the defs and split the `if`s in the definition of -`compareOfLessAndEq` -3. seeing if we can split by cases on the arguments, then see if the defs work themselves out - (useful when `compare` is defined via a `match` statement, as it is for `Bool`) -/ -macro "compareOfLessAndEq_rfl" : tactic => - `(tactic| (intros a b; first | rfl | - (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | - (induction a <;> induction b <;> simp +decide only))) - -/-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. -We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ -class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α where - /-- A linear order is total. -/ - le_total (a b : α) : a ≤ b ∨ b ≤ a - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableLE : DecidableRel (· ≤ · : α → α → Prop) - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE - /-- In a linearly ordered type, we assume the order relations are all decidable. -/ - decidableLT : DecidableRel (· < · : α → α → Prop) := - @decidableLTOfDecidableLE _ _ decidableLE - min := fun a b => if a ≤ b then a else b - max := fun a b => if a ≤ b then b else a - /-- The minimum function is equivalent to the one you get from `minOfLe`. -/ - min_def : ∀ a b, min a b = if a ≤ b then a else b := by intros; rfl - /-- The minimum function is equivalent to the one you get from `maxOfLe`. -/ - max_def : ∀ a b, max a b = if a ≤ b then b else a := by intros; rfl - compare a b := compareOfLessAndEq a b - /-- Comparison via `compare` is equal to the canonical comparison given decidable `<` and `=`. -/ - compare_eq_compareOfLessAndEq : ∀ a b, compare a b = compareOfLessAndEq a b := by - compareOfLessAndEq_rfl - -variable [LinearOrder α] {a b c : α} - -attribute [local instance] LinearOrder.decidableLE - -lemma le_total : ∀ a b : α, a ≤ b ∨ b ≤ a := LinearOrder.le_total - -lemma le_of_not_ge : ¬a ≥ b → a ≤ b := (le_total b a).resolve_left -lemma le_of_not_le : ¬a ≤ b → b ≤ a := (le_total a b).resolve_left -lemma lt_of_not_ge (h : ¬a ≥ b) : a < b := lt_of_le_not_le (le_of_not_ge h) h - -lemma lt_trichotomy (a b : α) : a < b ∨ a = b ∨ b < a := - Or.elim (le_total a b) - (fun h : a ≤ b => - Or.elim (Decidable.lt_or_eq_of_le h) (fun h : a < b => Or.inl h) fun h : a = b => - Or.inr (Or.inl h)) - fun h : b ≤ a => - Or.elim (Decidable.lt_or_eq_of_le h) (fun h : b < a => Or.inr (Or.inr h)) fun h : b = a => - Or.inr (Or.inl h.symm) - -lemma le_of_not_lt (h : ¬b < a) : a ≤ b := - match lt_trichotomy a b with - | Or.inl hlt => le_of_lt hlt - | Or.inr (Or.inl HEq) => HEq ▸ le_refl a - | Or.inr (Or.inr hgt) => absurd hgt h - -lemma le_of_not_gt : ¬a > b → a ≤ b := le_of_not_lt - -lemma lt_or_le (a b : α) : a < b ∨ b ≤ a := - if hba : b ≤ a then Or.inr hba else Or.inl <| lt_of_not_ge hba - -lemma le_or_lt (a b : α) : a ≤ b ∨ b < a := (lt_or_le b a).symm -lemma lt_or_ge : ∀ a b : α, a < b ∨ a ≥ b := lt_or_le -lemma le_or_gt : ∀ a b : α, a ≤ b ∨ a > b := le_or_lt - -lemma lt_or_gt_of_ne (h : a ≠ b) : a < b ∨ a > b := by simpa [h] using lt_trichotomy a b - -lemma ne_iff_lt_or_gt : a ≠ b ↔ a < b ∨ a > b := ⟨lt_or_gt_of_ne, (Or.elim · ne_of_lt ne_of_gt)⟩ - -lemma lt_iff_not_ge (x y : α) : x < y ↔ ¬x ≥ y := ⟨not_le_of_gt, lt_of_not_ge⟩ - -@[simp] lemma not_lt : ¬a < b ↔ b ≤ a := ⟨le_of_not_gt, not_lt_of_ge⟩ -@[simp] lemma not_le : ¬a ≤ b ↔ b < a := (lt_iff_not_ge _ _).symm - -instance (priority := 900) (a b : α) : Decidable (a < b) := LinearOrder.decidableLT a b -instance (priority := 900) (a b : α) : Decidable (a ≤ b) := LinearOrder.decidableLE a b -instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidableEq a b - -lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := - if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) - -/-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ -def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := - if h : x < y then h₁ h - else if h' : y < x then h₃ h' else h₂ (le_antisymm (le_of_not_gt h') (le_of_not_gt h)) - -namespace Nat - -/-! Deprecated properties of inequality on `Nat` -/ - -@[deprecated "No deprecation message was provided." (since := "2024-08-23")] -protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := - Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) - -set_option linter.deprecated false in -@[deprecated ltByCases (since := "2024-08-23")] -protected def ltByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : a = b → C) - (h₃ : b < a → C) : C := - Nat.ltGeByCases h₁ fun h₁ => Nat.ltGeByCases h₃ fun h => h₂ (Nat.le_antisymm h h₁) - -end Nat - -theorem le_imp_le_of_lt_imp_lt {α β} [Preorder α] [LinearOrder β] {a b : α} {c d : β} - (H : d < c → b < a) (h : a ≤ b) : c ≤ d := - le_of_not_lt fun h' => not_le_of_gt (H h') h - -lemma min_def (a b : α) : min a b = if a ≤ b then a else b := by rw [LinearOrder.min_def a] -lemma max_def (a b : α) : max a b = if a ≤ b then b else a := by rw [LinearOrder.max_def a] - --- Porting note: no `min_tac` tactic in the following series of lemmas - -lemma min_le_left (a b : α) : min a b ≤ a := by - if h : a ≤ b - then simp [min_def, if_pos h, le_refl] - else simp [min_def, if_neg h]; exact le_of_not_le h - -lemma min_le_right (a b : α) : min a b ≤ b := by - if h : a ≤ b - then simp [min_def, if_pos h]; exact h - else simp [min_def, if_neg h, le_refl] - -lemma le_min (h₁ : c ≤ a) (h₂ : c ≤ b) : c ≤ min a b := by - if h : a ≤ b - then simp [min_def, if_pos h]; exact h₁ - else simp [min_def, if_neg h]; exact h₂ - -lemma le_max_left (a b : α) : a ≤ max a b := by - if h : a ≤ b - then simp [max_def, if_pos h]; exact h - else simp [max_def, if_neg h, le_refl] - -lemma le_max_right (a b : α) : b ≤ max a b := by - if h : a ≤ b - then simp [max_def, if_pos h, le_refl] - else simp [max_def, if_neg h]; exact le_of_not_le h - -lemma max_le (h₁ : a ≤ c) (h₂ : b ≤ c) : max a b ≤ c := by - if h : a ≤ b - then simp [max_def, if_pos h]; exact h₂ - else simp [max_def, if_neg h]; exact h₁ - -lemma eq_min (h₁ : c ≤ a) (h₂ : c ≤ b) (h₃ : ∀ {d}, d ≤ a → d ≤ b → d ≤ c) : c = min a b := - le_antisymm (le_min h₁ h₂) (h₃ (min_le_left a b) (min_le_right a b)) - -lemma min_comm (a b : α) : min a b = min b a := - eq_min (min_le_right a b) (min_le_left a b) fun h₁ h₂ => le_min h₂ h₁ - -lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by - apply eq_min - · apply le_trans (min_le_left ..); apply min_le_left - · apply le_min - · apply le_trans (min_le_left ..); apply min_le_right - · apply min_le_right - · intro d h₁ h₂; apply le_min - · apply le_min h₁; apply le_trans h₂; apply min_le_left - · apply le_trans h₂; apply min_le_right - -lemma min_left_comm (a b c : α) : min a (min b c) = min b (min a c) := by - rw [← min_assoc, min_comm a, min_assoc] - -@[simp] lemma min_self (a : α) : min a a = a := by simp [min_def] - -lemma min_eq_left (h : a ≤ b) : min a b = a := by - apply Eq.symm; apply eq_min (le_refl _) h; intros; assumption - -lemma min_eq_right (h : b ≤ a) : min a b = b := min_comm b a ▸ min_eq_left h - -lemma eq_max (h₁ : a ≤ c) (h₂ : b ≤ c) (h₃ : ∀ {d}, a ≤ d → b ≤ d → c ≤ d) : - c = max a b := - le_antisymm (h₃ (le_max_left a b) (le_max_right a b)) (max_le h₁ h₂) - -lemma max_comm (a b : α) : max a b = max b a := - eq_max (le_max_right a b) (le_max_left a b) fun h₁ h₂ => max_le h₂ h₁ - -lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by - apply eq_max - · apply le_trans (le_max_left a b); apply le_max_left - · apply max_le - · apply le_trans (le_max_right a b); apply le_max_left - · apply le_max_right - · intro d h₁ h₂; apply max_le - · apply max_le h₁; apply le_trans (le_max_left _ _) h₂ - · apply le_trans (le_max_right _ _) h₂ - -lemma max_left_comm (a b c : α) : max a (max b c) = max b (max a c) := by - rw [← max_assoc, max_comm a, max_assoc] - -@[simp] lemma max_self (a : α) : max a a = a := by simp [max_def] - -lemma max_eq_left (h : b ≤ a) : max a b = a := by - apply Eq.symm; apply eq_max (le_refl _) h; intros; assumption - -lemma max_eq_right (h : a ≤ b) : max a b = b := max_comm b a ▸ max_eq_left h - -lemma min_eq_left_of_lt (h : a < b) : min a b = a := min_eq_left (le_of_lt h) -lemma min_eq_right_of_lt (h : b < a) : min a b = b := min_eq_right (le_of_lt h) -lemma max_eq_left_of_lt (h : b < a) : max a b = a := max_eq_left (le_of_lt h) -lemma max_eq_right_of_lt (h : a < b) : max a b = b := max_eq_right (le_of_lt h) - -lemma lt_min (h₁ : a < b) (h₂ : a < c) : a < min b c := by - cases le_total b c <;> simp [min_eq_left, min_eq_right, *] - -lemma max_lt (h₁ : a < c) (h₂ : b < c) : max a b < c := by - cases le_total a b <;> simp [max_eq_left, max_eq_right, *] - -section Ord - -lemma compare_lt_iff_lt : compare a b = .lt ↔ a < b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> simp only [*, lt_irrefl, reduceCtorEq] - -lemma compare_gt_iff_gt : compare a b = .gt ↔ a > b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> simp only [*, lt_irrefl, not_lt_of_gt, reduceCtorEq] - case _ h₁ h₂ => - have h : b < a := lt_trichotomy a b |>.resolve_left h₁ |>.resolve_left h₂ - rwa [true_iff] - -lemma compare_eq_iff_eq : compare a b = .eq ↔ a = b := by - rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] - split_ifs <;> try simp only [reduceCtorEq] - case _ h => rw [false_iff]; exact ne_iff_lt_or_gt.2 <| .inl h - case _ _ h => rwa [true_iff] - case _ _ h => rwa [false_iff] - -lemma compare_le_iff_le : compare a b ≠ .gt ↔ a ≤ b := by - cases h : compare a b <;> simp - · exact le_of_lt <| compare_lt_iff_lt.1 h - · exact le_of_eq <| compare_eq_iff_eq.1 h - · exact compare_gt_iff_gt.1 h - -lemma compare_ge_iff_ge : compare a b ≠ .lt ↔ a ≥ b := by - cases h : compare a b <;> simp - · exact compare_lt_iff_lt.1 h - · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h - · exact le_of_lt <| compare_gt_iff_gt.1 h - -lemma compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.Compares a b := by - cases o <;> simp only [Ordering.Compares] - · exact compare_lt_iff_lt - · exact compare_eq_iff_eq - · exact compare_gt_iff_gt - -theorem cmp_eq_compare (a b : α) : cmp a b = compare a b := by - refine ((compare_iff ..).2 ?_).symm - unfold cmp cmpUsing; split_ifs with h1 h2 - · exact h1 - · exact h2 - · exact le_antisymm (not_lt.1 h2) (not_lt.1 h1) - -theorem cmp_eq_compareOfLessAndEq (a b : α) : cmp a b = compareOfLessAndEq a b := - (cmp_eq_compare ..).trans (LinearOrder.compare_eq_compareOfLessAndEq ..) - -instance : Batteries.LawfulCmp (compare (α := α)) where - symm a b := by - cases h : compare a b <;> - simp only [Ordering.swap] <;> symm - · exact compare_gt_iff_gt.2 <| compare_lt_iff_lt.1 h - · exact compare_eq_iff_eq.2 <| compare_eq_iff_eq.1 h |>.symm - · exact compare_lt_iff_lt.2 <| compare_gt_iff_gt.1 h - le_trans := fun h₁ h₂ ↦ - compare_le_iff_le.2 <| le_trans (compare_le_iff_le.1 h₁) (compare_le_iff_le.1 h₂) - cmp_iff_beq := by simp [compare_eq_iff_eq] - cmp_iff_lt := by simp [compare_lt_iff_lt] - cmp_iff_le := by simp [compare_le_iff_le] - -end Ord - -end LinearOrder - -/-! ### Upper and lower sets -/ - -/-- An upper set in an order `α` is a set such that any element greater than one of its members is -also a member. Also called up-set, upward-closed set. -/ -def IsUpperSet {α : Type*} [LE α] (s : Set α) : Prop := - ∀ ⦃a b : α⦄, a ≤ b → a ∈ s → b ∈ s - -/-- A lower set in an order `α` is a set such that any element less than one of its members is also -a member. Also called down-set, downward-closed set. -/ -def IsLowerSet {α : Type*} [LE α] (s : Set α) : Prop := - ∀ ⦃a b : α⦄, b ≤ a → a ∈ s → b ∈ s - -@[inherit_doc IsUpperSet] -structure UpperSet (α : Type*) [LE α] where - /-- The carrier of an `UpperSet`. -/ - carrier : Set α - /-- The carrier of an `UpperSet` is an upper set. -/ - upper' : IsUpperSet carrier - -extend_docs UpperSet before "The type of upper sets of an order." - -@[inherit_doc IsLowerSet] -structure LowerSet (α : Type*) [LE α] where - /-- The carrier of a `LowerSet`. -/ - carrier : Set α - /-- The carrier of a `LowerSet` is a lower set. -/ - lower' : IsLowerSet carrier - -extend_docs LowerSet before "The type of lower sets of an order." diff --git a/Mathlib/Order/Defs/LinearOrder.lean b/Mathlib/Order/Defs/LinearOrder.lean new file mode 100644 index 0000000000000..ec281a950f1d6 --- /dev/null +++ b/Mathlib/Order/Defs/LinearOrder.lean @@ -0,0 +1,311 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Batteries.Classes.Order +import Batteries.Tactic.Trans +import Mathlib.Data.Ordering.Basic +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar +import Mathlib.Order.Defs.PartialOrder + +/-! +# Orders + +Defines classes for linear orders and proves some basic lemmas about them. +-/ + +variable {α : Type*} + +section LinearOrder + +/-! +### Definition of `LinearOrder` and lemmas about types with a linear order +-/ + +/-- Default definition of `max`. -/ +def maxDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := + if a ≤ b then b else a + +/-- Default definition of `min`. -/ +def minDefault [LE α] [DecidableRel ((· ≤ ·) : α → α → Prop)] (a b : α) := + if a ≤ b then a else b + +/-- This attempts to prove that a given instance of `compare` is equal to `compareOfLessAndEq` by +introducing the arguments and trying the following approaches in order: + +1. seeing if `rfl` works +2. seeing if the `compare` at hand is nonetheless essentially `compareOfLessAndEq`, but, because of +implicit arguments, requires us to unfold the defs and split the `if`s in the definition of +`compareOfLessAndEq` +3. seeing if we can split by cases on the arguments, then see if the defs work themselves out + (useful when `compare` is defined via a `match` statement, as it is for `Bool`) -/ +macro "compareOfLessAndEq_rfl" : tactic => + `(tactic| (intros a b; first | rfl | + (simp only [compare, compareOfLessAndEq]; split_ifs <;> rfl) | + (induction a <;> induction b <;> simp +decide only))) + +/-- A linear order is reflexive, transitive, antisymmetric and total relation `≤`. +We assume that every linear ordered type has decidable `(≤)`, `(<)`, and `(=)`. -/ +class LinearOrder (α : Type*) extends PartialOrder α, Min α, Max α, Ord α where + /-- A linear order is total. -/ + le_total (a b : α) : a ≤ b ∨ b ≤ a + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableLE : DecidableRel (· ≤ · : α → α → Prop) + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableEq : DecidableEq α := @decidableEqOfDecidableLE _ _ decidableLE + /-- In a linearly ordered type, we assume the order relations are all decidable. -/ + decidableLT : DecidableRel (· < · : α → α → Prop) := + @decidableLTOfDecidableLE _ _ decidableLE + min := fun a b => if a ≤ b then a else b + max := fun a b => if a ≤ b then b else a + /-- The minimum function is equivalent to the one you get from `minOfLe`. -/ + min_def : ∀ a b, min a b = if a ≤ b then a else b := by intros; rfl + /-- The minimum function is equivalent to the one you get from `maxOfLe`. -/ + max_def : ∀ a b, max a b = if a ≤ b then b else a := by intros; rfl + compare a b := compareOfLessAndEq a b + /-- Comparison via `compare` is equal to the canonical comparison given decidable `<` and `=`. -/ + compare_eq_compareOfLessAndEq : ∀ a b, compare a b = compareOfLessAndEq a b := by + compareOfLessAndEq_rfl + +variable [LinearOrder α] {a b c : α} + +attribute [local instance] LinearOrder.decidableLE + +lemma le_total : ∀ a b : α, a ≤ b ∨ b ≤ a := LinearOrder.le_total + +lemma le_of_not_ge : ¬a ≥ b → a ≤ b := (le_total b a).resolve_left +lemma le_of_not_le : ¬a ≤ b → b ≤ a := (le_total a b).resolve_left +lemma lt_of_not_ge (h : ¬a ≥ b) : a < b := lt_of_le_not_le (le_of_not_ge h) h + +lemma lt_trichotomy (a b : α) : a < b ∨ a = b ∨ b < a := + Or.elim (le_total a b) + (fun h : a ≤ b => + Or.elim (Decidable.lt_or_eq_of_le h) (fun h : a < b => Or.inl h) fun h : a = b => + Or.inr (Or.inl h)) + fun h : b ≤ a => + Or.elim (Decidable.lt_or_eq_of_le h) (fun h : b < a => Or.inr (Or.inr h)) fun h : b = a => + Or.inr (Or.inl h.symm) + +lemma le_of_not_lt (h : ¬b < a) : a ≤ b := + match lt_trichotomy a b with + | Or.inl hlt => le_of_lt hlt + | Or.inr (Or.inl HEq) => HEq ▸ le_refl a + | Or.inr (Or.inr hgt) => absurd hgt h + +lemma le_of_not_gt : ¬a > b → a ≤ b := le_of_not_lt + +lemma lt_or_le (a b : α) : a < b ∨ b ≤ a := + if hba : b ≤ a then Or.inr hba else Or.inl <| lt_of_not_ge hba + +lemma le_or_lt (a b : α) : a ≤ b ∨ b < a := (lt_or_le b a).symm +lemma lt_or_ge : ∀ a b : α, a < b ∨ a ≥ b := lt_or_le +lemma le_or_gt : ∀ a b : α, a ≤ b ∨ a > b := le_or_lt + +lemma lt_or_gt_of_ne (h : a ≠ b) : a < b ∨ a > b := by simpa [h] using lt_trichotomy a b + +lemma ne_iff_lt_or_gt : a ≠ b ↔ a < b ∨ a > b := ⟨lt_or_gt_of_ne, (Or.elim · ne_of_lt ne_of_gt)⟩ + +lemma lt_iff_not_ge (x y : α) : x < y ↔ ¬x ≥ y := ⟨not_le_of_gt, lt_of_not_ge⟩ + +@[simp] lemma not_lt : ¬a < b ↔ b ≤ a := ⟨le_of_not_gt, not_lt_of_ge⟩ +@[simp] lemma not_le : ¬a ≤ b ↔ b < a := (lt_iff_not_ge _ _).symm + +instance (priority := 900) (a b : α) : Decidable (a < b) := LinearOrder.decidableLT a b +instance (priority := 900) (a b : α) : Decidable (a ≤ b) := LinearOrder.decidableLE a b +instance (priority := 900) (a b : α) : Decidable (a = b) := LinearOrder.decidableEq a b + +lemma eq_or_lt_of_not_lt (h : ¬a < b) : a = b ∨ b < a := + if h₁ : a = b then Or.inl h₁ else Or.inr (lt_of_not_ge fun hge => h (lt_of_le_of_ne hge h₁)) + +/-- Perform a case-split on the ordering of `x` and `y` in a decidable linear order. -/ +def ltByCases (x y : α) {P : Sort*} (h₁ : x < y → P) (h₂ : x = y → P) (h₃ : y < x → P) : P := + if h : x < y then h₁ h + else if h' : y < x then h₃ h' else h₂ (le_antisymm (le_of_not_gt h') (le_of_not_gt h)) + +namespace Nat + +/-! Deprecated properties of inequality on `Nat` -/ + +@[deprecated "No deprecation message was provided." (since := "2024-08-23")] +protected def ltGeByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : b ≤ a → C) : C := + Decidable.byCases h₁ fun h => h₂ (Or.elim (Nat.lt_or_ge a b) (fun a => absurd a h) fun a => a) + +set_option linter.deprecated false in +@[deprecated ltByCases (since := "2024-08-23")] +protected def ltByCases {a b : Nat} {C : Sort*} (h₁ : a < b → C) (h₂ : a = b → C) + (h₃ : b < a → C) : C := + Nat.ltGeByCases h₁ fun h₁ => Nat.ltGeByCases h₃ fun h => h₂ (Nat.le_antisymm h h₁) + +end Nat + +theorem le_imp_le_of_lt_imp_lt {α β} [Preorder α] [LinearOrder β] {a b : α} {c d : β} + (H : d < c → b < a) (h : a ≤ b) : c ≤ d := + le_of_not_lt fun h' => not_le_of_gt (H h') h + +lemma min_def (a b : α) : min a b = if a ≤ b then a else b := by rw [LinearOrder.min_def a] +lemma max_def (a b : α) : max a b = if a ≤ b then b else a := by rw [LinearOrder.max_def a] + +-- Porting note: no `min_tac` tactic in the following series of lemmas + +lemma min_le_left (a b : α) : min a b ≤ a := by + if h : a ≤ b + then simp [min_def, if_pos h, le_refl] + else simp [min_def, if_neg h]; exact le_of_not_le h + +lemma min_le_right (a b : α) : min a b ≤ b := by + if h : a ≤ b + then simp [min_def, if_pos h]; exact h + else simp [min_def, if_neg h, le_refl] + +lemma le_min (h₁ : c ≤ a) (h₂ : c ≤ b) : c ≤ min a b := by + if h : a ≤ b + then simp [min_def, if_pos h]; exact h₁ + else simp [min_def, if_neg h]; exact h₂ + +lemma le_max_left (a b : α) : a ≤ max a b := by + if h : a ≤ b + then simp [max_def, if_pos h]; exact h + else simp [max_def, if_neg h, le_refl] + +lemma le_max_right (a b : α) : b ≤ max a b := by + if h : a ≤ b + then simp [max_def, if_pos h, le_refl] + else simp [max_def, if_neg h]; exact le_of_not_le h + +lemma max_le (h₁ : a ≤ c) (h₂ : b ≤ c) : max a b ≤ c := by + if h : a ≤ b + then simp [max_def, if_pos h]; exact h₂ + else simp [max_def, if_neg h]; exact h₁ + +lemma eq_min (h₁ : c ≤ a) (h₂ : c ≤ b) (h₃ : ∀ {d}, d ≤ a → d ≤ b → d ≤ c) : c = min a b := + le_antisymm (le_min h₁ h₂) (h₃ (min_le_left a b) (min_le_right a b)) + +lemma min_comm (a b : α) : min a b = min b a := + eq_min (min_le_right a b) (min_le_left a b) fun h₁ h₂ => le_min h₂ h₁ + +lemma min_assoc (a b c : α) : min (min a b) c = min a (min b c) := by + apply eq_min + · apply le_trans (min_le_left ..); apply min_le_left + · apply le_min + · apply le_trans (min_le_left ..); apply min_le_right + · apply min_le_right + · intro d h₁ h₂; apply le_min + · apply le_min h₁; apply le_trans h₂; apply min_le_left + · apply le_trans h₂; apply min_le_right + +lemma min_left_comm (a b c : α) : min a (min b c) = min b (min a c) := by + rw [← min_assoc, min_comm a, min_assoc] + +@[simp] lemma min_self (a : α) : min a a = a := by simp [min_def] + +lemma min_eq_left (h : a ≤ b) : min a b = a := by + apply Eq.symm; apply eq_min (le_refl _) h; intros; assumption + +lemma min_eq_right (h : b ≤ a) : min a b = b := min_comm b a ▸ min_eq_left h + +lemma eq_max (h₁ : a ≤ c) (h₂ : b ≤ c) (h₃ : ∀ {d}, a ≤ d → b ≤ d → c ≤ d) : + c = max a b := + le_antisymm (h₃ (le_max_left a b) (le_max_right a b)) (max_le h₁ h₂) + +lemma max_comm (a b : α) : max a b = max b a := + eq_max (le_max_right a b) (le_max_left a b) fun h₁ h₂ => max_le h₂ h₁ + +lemma max_assoc (a b c : α) : max (max a b) c = max a (max b c) := by + apply eq_max + · apply le_trans (le_max_left a b); apply le_max_left + · apply max_le + · apply le_trans (le_max_right a b); apply le_max_left + · apply le_max_right + · intro d h₁ h₂; apply max_le + · apply max_le h₁; apply le_trans (le_max_left _ _) h₂ + · apply le_trans (le_max_right _ _) h₂ + +lemma max_left_comm (a b c : α) : max a (max b c) = max b (max a c) := by + rw [← max_assoc, max_comm a, max_assoc] + +@[simp] lemma max_self (a : α) : max a a = a := by simp [max_def] + +lemma max_eq_left (h : b ≤ a) : max a b = a := by + apply Eq.symm; apply eq_max (le_refl _) h; intros; assumption + +lemma max_eq_right (h : a ≤ b) : max a b = b := max_comm b a ▸ max_eq_left h + +lemma min_eq_left_of_lt (h : a < b) : min a b = a := min_eq_left (le_of_lt h) +lemma min_eq_right_of_lt (h : b < a) : min a b = b := min_eq_right (le_of_lt h) +lemma max_eq_left_of_lt (h : b < a) : max a b = a := max_eq_left (le_of_lt h) +lemma max_eq_right_of_lt (h : a < b) : max a b = b := max_eq_right (le_of_lt h) + +lemma lt_min (h₁ : a < b) (h₂ : a < c) : a < min b c := by + cases le_total b c <;> simp [min_eq_left, min_eq_right, *] + +lemma max_lt (h₁ : a < c) (h₂ : b < c) : max a b < c := by + cases le_total a b <;> simp [max_eq_left, max_eq_right, *] + +section Ord + +lemma compare_lt_iff_lt : compare a b = .lt ↔ a < b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> simp only [*, lt_irrefl, reduceCtorEq] + +lemma compare_gt_iff_gt : compare a b = .gt ↔ a > b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> simp only [*, lt_irrefl, not_lt_of_gt, reduceCtorEq] + case _ h₁ h₂ => + have h : b < a := lt_trichotomy a b |>.resolve_left h₁ |>.resolve_left h₂ + rwa [true_iff] + +lemma compare_eq_iff_eq : compare a b = .eq ↔ a = b := by + rw [LinearOrder.compare_eq_compareOfLessAndEq, compareOfLessAndEq] + split_ifs <;> try simp only [reduceCtorEq] + case _ h => rw [false_iff]; exact ne_iff_lt_or_gt.2 <| .inl h + case _ _ h => rwa [true_iff] + case _ _ h => rwa [false_iff] + +lemma compare_le_iff_le : compare a b ≠ .gt ↔ a ≤ b := by + cases h : compare a b <;> simp + · exact le_of_lt <| compare_lt_iff_lt.1 h + · exact le_of_eq <| compare_eq_iff_eq.1 h + · exact compare_gt_iff_gt.1 h + +lemma compare_ge_iff_ge : compare a b ≠ .lt ↔ a ≥ b := by + cases h : compare a b <;> simp + · exact compare_lt_iff_lt.1 h + · exact le_of_eq <| (·.symm) <| compare_eq_iff_eq.1 h + · exact le_of_lt <| compare_gt_iff_gt.1 h + +lemma compare_iff (a b : α) {o : Ordering} : compare a b = o ↔ o.Compares a b := by + cases o <;> simp only [Ordering.Compares] + · exact compare_lt_iff_lt + · exact compare_eq_iff_eq + · exact compare_gt_iff_gt + +theorem cmp_eq_compare (a b : α) : cmp a b = compare a b := by + refine ((compare_iff ..).2 ?_).symm + unfold cmp cmpUsing; split_ifs with h1 h2 + · exact h1 + · exact h2 + · exact le_antisymm (not_lt.1 h2) (not_lt.1 h1) + +theorem cmp_eq_compareOfLessAndEq (a b : α) : cmp a b = compareOfLessAndEq a b := + (cmp_eq_compare ..).trans (LinearOrder.compare_eq_compareOfLessAndEq ..) + +instance : Batteries.LawfulCmp (compare (α := α)) where + symm a b := by + cases h : compare a b <;> + simp only [Ordering.swap] <;> symm + · exact compare_gt_iff_gt.2 <| compare_lt_iff_lt.1 h + · exact compare_eq_iff_eq.2 <| compare_eq_iff_eq.1 h |>.symm + · exact compare_lt_iff_lt.2 <| compare_gt_iff_gt.1 h + le_trans := fun h₁ h₂ ↦ + compare_le_iff_le.2 <| le_trans (compare_le_iff_le.1 h₁) (compare_le_iff_le.1 h₂) + cmp_iff_beq := by simp [compare_eq_iff_eq] + cmp_iff_lt := by simp [compare_lt_iff_lt] + cmp_iff_le := by simp [compare_le_iff_le] + +end Ord + +end LinearOrder diff --git a/Mathlib/Order/Defs/PartialOrder.lean b/Mathlib/Order/Defs/PartialOrder.lean new file mode 100644 index 0000000000000..377677f6d9855 --- /dev/null +++ b/Mathlib/Order/Defs/PartialOrder.lean @@ -0,0 +1,158 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Batteries.Tactic.Trans +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar + +/-! +# Orders + +Defines classes for preorders and partial orders +and proves some basic lemmas about them. +-/ + +variable {α : Type*} + +section Preorder + +/-! +### Definition of `Preorder` and lemmas about types with a `Preorder` +-/ + +/-- A preorder is a reflexive, transitive relation `≤` with `a < b` defined in the obvious way. -/ +class Preorder (α : Type*) extends LE α, LT α where + le_refl : ∀ a : α, a ≤ a + le_trans : ∀ a b c : α, a ≤ b → b ≤ c → a ≤ c + lt := fun a b => a ≤ b ∧ ¬b ≤ a + lt_iff_le_not_le : ∀ a b : α, a < b ↔ a ≤ b ∧ ¬b ≤ a := by intros; rfl + +variable [Preorder α] {a b c : α} + +/-- The relation `≤` on a preorder is reflexive. -/ +@[refl, simp] lemma le_refl : ∀ a : α, a ≤ a := Preorder.le_refl + +/-- A version of `le_refl` where the argument is implicit -/ +lemma le_rfl : a ≤ a := le_refl a + +/-- The relation `≤` on a preorder is transitive. -/ +@[trans] lemma le_trans : a ≤ b → b ≤ c → a ≤ c := Preorder.le_trans _ _ _ + +lemma lt_iff_le_not_le : a < b ↔ a ≤ b ∧ ¬b ≤ a := Preorder.lt_iff_le_not_le _ _ + +lemma lt_of_le_not_le (hab : a ≤ b) (hba : ¬ b ≤ a) : a < b := lt_iff_le_not_le.2 ⟨hab, hba⟩ + +@[deprecated "No deprecation message was provided." (since := "2024-07-30")] +theorem le_not_le_of_lt : ∀ {a b : α}, a < b → a ≤ b ∧ ¬b ≤ a + | _a, _b, hab => lt_iff_le_not_le.mp hab + +lemma le_of_eq (hab : a = b) : a ≤ b := by rw [hab] +lemma le_of_lt (hab : a < b) : a ≤ b := (lt_iff_le_not_le.1 hab).1 +lemma not_le_of_lt (hab : a < b) : ¬ b ≤ a := (lt_iff_le_not_le.1 hab).2 +lemma not_le_of_gt (hab : a > b) : ¬a ≤ b := not_le_of_lt hab +lemma not_lt_of_le (hab : a ≤ b) : ¬ b < a := imp_not_comm.1 not_le_of_lt hab +lemma not_lt_of_ge (hab : a ≥ b) : ¬a < b := not_lt_of_le hab + +alias LT.lt.not_le := not_le_of_lt +alias LE.le.not_lt := not_lt_of_le + +@[trans] lemma ge_trans : a ≥ b → b ≥ c → a ≥ c := fun h₁ h₂ => le_trans h₂ h₁ + +lemma lt_irrefl (a : α) : ¬a < a := fun h ↦ not_le_of_lt h le_rfl +lemma gt_irrefl (a : α) : ¬a > a := lt_irrefl _ + +@[trans] lemma lt_of_lt_of_le (hab : a < b) (hbc : b ≤ c) : a < c := + lt_of_le_not_le (le_trans (le_of_lt hab) hbc) fun hca ↦ not_le_of_lt hab (le_trans hbc hca) + +@[trans] lemma lt_of_le_of_lt (hab : a ≤ b) (hbc : b < c) : a < c := + lt_of_le_not_le (le_trans hab (le_of_lt hbc)) fun hca ↦ not_le_of_lt hbc (le_trans hca hab) + +@[trans] lemma gt_of_gt_of_ge (h₁ : a > b) (h₂ : b ≥ c) : a > c := lt_of_le_of_lt h₂ h₁ +@[trans] lemma gt_of_ge_of_gt (h₁ : a ≥ b) (h₂ : b > c) : a > c := lt_of_lt_of_le h₂ h₁ + +@[trans] lemma lt_trans (hab : a < b) (hbc : b < c) : a < c := lt_of_lt_of_le hab (le_of_lt hbc) +@[trans] lemma gt_trans : a > b → b > c → a > c := fun h₁ h₂ => lt_trans h₂ h₁ + +lemma ne_of_lt (h : a < b) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) +lemma ne_of_gt (h : b < a) : a ≠ b := fun he => absurd h (he ▸ lt_irrefl a) +lemma lt_asymm (h : a < b) : ¬b < a := fun h1 : b < a => lt_irrefl a (lt_trans h h1) + +alias not_lt_of_gt := lt_asymm +alias not_lt_of_lt := lt_asymm + +lemma le_of_lt_or_eq (h : a < b ∨ a = b) : a ≤ b := h.elim le_of_lt le_of_eq +lemma le_of_eq_or_lt (h : a = b ∨ a < b) : a ≤ b := h.elim le_of_eq le_of_lt + +instance (priority := 900) : @Trans α α α LE.le LE.le LE.le := ⟨le_trans⟩ +instance (priority := 900) : @Trans α α α LT.lt LT.lt LT.lt := ⟨lt_trans⟩ +instance (priority := 900) : @Trans α α α LT.lt LE.le LT.lt := ⟨lt_of_lt_of_le⟩ +instance (priority := 900) : @Trans α α α LE.le LT.lt LT.lt := ⟨lt_of_le_of_lt⟩ +instance (priority := 900) : @Trans α α α GE.ge GE.ge GE.ge := ⟨ge_trans⟩ +instance (priority := 900) : @Trans α α α GT.gt GT.gt GT.gt := ⟨gt_trans⟩ +instance (priority := 900) : @Trans α α α GT.gt GE.ge GT.gt := ⟨gt_of_gt_of_ge⟩ +instance (priority := 900) : @Trans α α α GE.ge GT.gt GT.gt := ⟨gt_of_ge_of_gt⟩ + +/-- `<` is decidable if `≤` is. -/ +def decidableLTOfDecidableLE [@DecidableRel α (· ≤ ·)] : @DecidableRel α (· < ·) + | a, b => + if hab : a ≤ b then + if hba : b ≤ a then isFalse fun hab' => not_le_of_gt hab' hba + else isTrue <| lt_of_le_not_le hab hba + else isFalse fun hab' => hab (le_of_lt hab') + +end Preorder + +section PartialOrder + +/-! +### Definition of `PartialOrder` and lemmas about types with a partial order +-/ + +/-- A partial order is a reflexive, transitive, antisymmetric relation `≤`. -/ +class PartialOrder (α : Type*) extends Preorder α where + le_antisymm : ∀ a b : α, a ≤ b → b ≤ a → a = b + +variable [PartialOrder α] {a b : α} + +lemma le_antisymm : a ≤ b → b ≤ a → a = b := PartialOrder.le_antisymm _ _ + +alias eq_of_le_of_le := le_antisymm + +lemma le_antisymm_iff : a = b ↔ a ≤ b ∧ b ≤ a := + ⟨fun e => ⟨le_of_eq e, le_of_eq e.symm⟩, fun ⟨h1, h2⟩ => le_antisymm h1 h2⟩ + +lemma lt_of_le_of_ne : a ≤ b → a ≠ b → a < b := fun h₁ h₂ => + lt_of_le_not_le h₁ <| mt (le_antisymm h₁) h₂ + +/-- Equality is decidable if `≤` is. -/ +def decidableEqOfDecidableLE [@DecidableRel α (· ≤ ·)] : DecidableEq α + | a, b => + if hab : a ≤ b then + if hba : b ≤ a then isTrue (le_antisymm hab hba) else isFalse fun heq => hba (heq ▸ le_refl _) + else isFalse fun heq => hab (heq ▸ le_refl _) + +namespace Decidable + +variable [@DecidableRel α (· ≤ ·)] + +lemma lt_or_eq_of_le (hab : a ≤ b) : a < b ∨ a = b := + if hba : b ≤ a then Or.inr (le_antisymm hab hba) else Or.inl (lt_of_le_not_le hab hba) + +lemma eq_or_lt_of_le (hab : a ≤ b) : a = b ∨ a < b := + (lt_or_eq_of_le hab).symm + +lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := + ⟨lt_or_eq_of_le, le_of_lt_or_eq⟩ + +end Decidable + +attribute [local instance] Classical.propDecidable + +lemma lt_or_eq_of_le : a ≤ b → a < b ∨ a = b := Decidable.lt_or_eq_of_le +lemma le_iff_lt_or_eq : a ≤ b ↔ a < b ∨ a = b := Decidable.le_iff_lt_or_eq + +end PartialOrder diff --git a/Mathlib/Order/Defs/Unbundled.lean b/Mathlib/Order/Defs/Unbundled.lean new file mode 100644 index 0000000000000..c666d5f19cd3b --- /dev/null +++ b/Mathlib/Order/Defs/Unbundled.lean @@ -0,0 +1,278 @@ +/- +Copyright (c) 2016 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura +-/ +import Mathlib.Data.Set.Defs +import Mathlib.Tactic.ExtendDoc +import Mathlib.Tactic.Lemma +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.TypeStar + +/-! +# Orders + +Defines classes for preorders, partial orders, and linear orders +and proves some basic lemmas about them. +-/ + +/-! ### Unbundled classes -/ + +/-- An empty relation does not relate any elements. -/ +@[nolint unusedArguments] def EmptyRelation {α : Sort*} := fun _ _ : α ↦ False + +/-- `IsIrrefl X r` means the binary relation `r` on `X` is irreflexive (that is, `r x x` never +holds). -/ +class IsIrrefl (α : Sort*) (r : α → α → Prop) : Prop where + irrefl : ∀ a, ¬r a a + +/-- `IsRefl X r` means the binary relation `r` on `X` is reflexive. -/ +class IsRefl (α : Sort*) (r : α → α → Prop) : Prop where + refl : ∀ a, r a a + +/-- `IsSymm X r` means the binary relation `r` on `X` is symmetric. -/ +class IsSymm (α : Sort*) (r : α → α → Prop) : Prop where + symm : ∀ a b, r a b → r b a + +/-- `IsAsymm X r` means that the binary relation `r` on `X` is asymmetric, that is, +`r a b → ¬ r b a`. -/ +class IsAsymm (α : Sort*) (r : α → α → Prop) : Prop where + asymm : ∀ a b, r a b → ¬r b a + +/-- `IsAntisymm X r` means the binary relation `r` on `X` is antisymmetric. -/ +class IsAntisymm (α : Sort*) (r : α → α → Prop) : Prop where + antisymm : ∀ a b, r a b → r b a → a = b + +instance (priority := 100) IsAsymm.toIsAntisymm {α : Sort*} (r : α → α → Prop) [IsAsymm α r] : + IsAntisymm α r where + antisymm _ _ hx hy := (IsAsymm.asymm _ _ hx hy).elim + +/-- `IsTrans X r` means the binary relation `r` on `X` is transitive. -/ +class IsTrans (α : Sort*) (r : α → α → Prop) : Prop where + trans : ∀ a b c, r a b → r b c → r a c + +instance {α : Sort*} {r : α → α → Prop} [IsTrans α r] : Trans r r r := + ⟨IsTrans.trans _ _ _⟩ + +instance (priority := 100) {α : Sort*} {r : α → α → Prop} [Trans r r r] : IsTrans α r := + ⟨fun _ _ _ => Trans.trans⟩ + +/-- `IsTotal X r` means that the binary relation `r` on `X` is total, that is, that for any +`x y : X` we have `r x y` or `r y x`. -/ +class IsTotal (α : Sort*) (r : α → α → Prop) : Prop where + total : ∀ a b, r a b ∨ r b a + +/-- `IsPreorder X r` means that the binary relation `r` on `X` is a pre-order, that is, reflexive +and transitive. -/ +class IsPreorder (α : Sort*) (r : α → α → Prop) extends IsRefl α r, IsTrans α r : Prop + +/-- `IsPartialOrder X r` means that the binary relation `r` on `X` is a partial order, that is, +`IsPreorder X r` and `IsAntisymm X r`. -/ +class IsPartialOrder (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsAntisymm α r : Prop + +/-- `IsLinearOrder X r` means that the binary relation `r` on `X` is a linear order, that is, +`IsPartialOrder X r` and `IsTotal X r`. -/ +class IsLinearOrder (α : Sort*) (r : α → α → Prop) extends IsPartialOrder α r, IsTotal α r : Prop + +/-- `IsEquiv X r` means that the binary relation `r` on `X` is an equivalence relation, that +is, `IsPreorder X r` and `IsSymm X r`. -/ +class IsEquiv (α : Sort*) (r : α → α → Prop) extends IsPreorder α r, IsSymm α r : Prop + +/-- `IsStrictOrder X r` means that the binary relation `r` on `X` is a strict order, that is, +`IsIrrefl X r` and `IsTrans X r`. -/ +class IsStrictOrder (α : Sort*) (r : α → α → Prop) extends IsIrrefl α r, IsTrans α r : Prop + +/-- `IsStrictWeakOrder X lt` means that the binary relation `lt` on `X` is a strict weak order, +that is, `IsStrictOrder X lt` and `¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a`. -/ +class IsStrictWeakOrder (α : Sort*) (lt : α → α → Prop) extends IsStrictOrder α lt : Prop where + incomp_trans : ∀ a b c, ¬lt a b ∧ ¬lt b a → ¬lt b c ∧ ¬lt c b → ¬lt a c ∧ ¬lt c a + +/-- `IsTrichotomous X lt` means that the binary relation `lt` on `X` is trichotomous, that is, +either `lt a b` or `a = b` or `lt b a` for any `a` and `b`. -/ +class IsTrichotomous (α : Sort*) (lt : α → α → Prop) : Prop where + trichotomous : ∀ a b, lt a b ∨ a = b ∨ lt b a + +/-- `IsStrictTotalOrder X lt` means that the binary relation `lt` on `X` is a strict total order, +that is, `IsTrichotomous X lt` and `IsStrictOrder X lt`. -/ +class IsStrictTotalOrder (α : Sort*) (lt : α → α → Prop) extends IsTrichotomous α lt, + IsStrictOrder α lt : Prop + +/-- Equality is an equivalence relation. -/ +instance eq_isEquiv (α : Sort*) : IsEquiv α (· = ·) where + symm := @Eq.symm _ + trans := @Eq.trans _ + refl := Eq.refl + +/-- `Iff` is an equivalence relation. -/ +instance iff_isEquiv : IsEquiv Prop Iff where + symm := @Iff.symm + trans := @Iff.trans + refl := @Iff.refl + +section + +variable {α : Sort*} {r : α → α → Prop} {a b c : α} + +/-- Local notation for an arbitrary binary relation `r`. -/ +local infixl:50 " ≺ " => r + +lemma irrefl [IsIrrefl α r] (a : α) : ¬a ≺ a := IsIrrefl.irrefl a +lemma refl [IsRefl α r] (a : α) : a ≺ a := IsRefl.refl a +lemma trans [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := IsTrans.trans _ _ _ +lemma symm [IsSymm α r] : a ≺ b → b ≺ a := IsSymm.symm _ _ +lemma antisymm [IsAntisymm α r] : a ≺ b → b ≺ a → a = b := IsAntisymm.antisymm _ _ +lemma asymm [IsAsymm α r] : a ≺ b → ¬b ≺ a := IsAsymm.asymm _ _ + +lemma trichotomous [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := + IsTrichotomous.trichotomous + +instance (priority := 90) isAsymm_of_isTrans_of_isIrrefl [IsTrans α r] [IsIrrefl α r] : + IsAsymm α r := + ⟨fun a _b h₁ h₂ => absurd (_root_.trans h₁ h₂) (irrefl a)⟩ + +instance IsIrrefl.decide [DecidableRel r] [IsIrrefl α r] : + IsIrrefl α (fun a b => decide (r a b) = true) where + irrefl := fun a => by simpa using irrefl a + +instance IsRefl.decide [DecidableRel r] [IsRefl α r] : + IsRefl α (fun a b => decide (r a b) = true) where + refl := fun a => by simpa using refl a + +instance IsTrans.decide [DecidableRel r] [IsTrans α r] : + IsTrans α (fun a b => decide (r a b) = true) where + trans := fun a b c => by simpa using trans a b c + +instance IsSymm.decide [DecidableRel r] [IsSymm α r] : + IsSymm α (fun a b => decide (r a b) = true) where + symm := fun a b => by simpa using symm a b + +instance IsAntisymm.decide [DecidableRel r] [IsAntisymm α r] : + IsAntisymm α (fun a b => decide (r a b) = true) where + antisymm := fun a b h₁ h₂ => antisymm _ _ (by simpa using h₁) (by simpa using h₂) + +instance IsAsymm.decide [DecidableRel r] [IsAsymm α r] : + IsAsymm α (fun a b => decide (r a b) = true) where + asymm := fun a b => by simpa using asymm a b + +instance IsTotal.decide [DecidableRel r] [IsTotal α r] : + IsTotal α (fun a b => decide (r a b) = true) where + total := fun a b => by simpa using total a b + +instance IsTrichotomous.decide [DecidableRel r] [IsTrichotomous α r] : + IsTrichotomous α (fun a b => decide (r a b) = true) where + trichotomous := fun a b => by simpa using trichotomous a b + +variable (r) + +@[elab_without_expected_type] lemma irrefl_of [IsIrrefl α r] (a : α) : ¬a ≺ a := irrefl a +@[elab_without_expected_type] lemma refl_of [IsRefl α r] (a : α) : a ≺ a := refl a +@[elab_without_expected_type] lemma trans_of [IsTrans α r] : a ≺ b → b ≺ c → a ≺ c := _root_.trans +@[elab_without_expected_type] lemma symm_of [IsSymm α r] : a ≺ b → b ≺ a := symm +@[elab_without_expected_type] lemma asymm_of [IsAsymm α r] : a ≺ b → ¬b ≺ a := asymm + +@[elab_without_expected_type] +lemma total_of [IsTotal α r] (a b : α) : a ≺ b ∨ b ≺ a := IsTotal.total _ _ + +@[elab_without_expected_type] +lemma trichotomous_of [IsTrichotomous α r] : ∀ a b : α, a ≺ b ∨ a = b ∨ b ≺ a := trichotomous + +section + +/-- `IsRefl` as a definition, suitable for use in proofs. -/ +def Reflexive := ∀ x, x ≺ x + +/-- `IsSymm` as a definition, suitable for use in proofs. -/ +def Symmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x + +/-- `IsTrans` as a definition, suitable for use in proofs. -/ +def Transitive := ∀ ⦃x y z⦄, x ≺ y → y ≺ z → x ≺ z + +/-- `IsIrrefl` as a definition, suitable for use in proofs. -/ +def Irreflexive := ∀ x, ¬x ≺ x + +/-- `IsAntisymm` as a definition, suitable for use in proofs. -/ +def AntiSymmetric := ∀ ⦃x y⦄, x ≺ y → y ≺ x → x = y + +/-- `IsTotal` as a definition, suitable for use in proofs. -/ +def Total := ∀ x y, x ≺ y ∨ y ≺ x + +@[deprecated Equivalence.refl (since := "2024-09-13")] +theorem Equivalence.reflexive (h : Equivalence r) : Reflexive r := h.refl + +@[deprecated Equivalence.symm (since := "2024-09-13")] +theorem Equivalence.symmetric (h : Equivalence r) : Symmetric r := + fun _ _ ↦ h.symm + +@[deprecated Equivalence.trans (since := "2024-09-13")] +theorem Equivalence.transitive (h : Equivalence r) : Transitive r := + fun _ _ _ ↦ h.trans + +variable {β : Sort*} (r : β → β → Prop) (f : α → β) + +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] +theorem InvImage.trans (h : Transitive r) : Transitive (InvImage r f) := + fun (a₁ a₂ a₃ : α) (h₁ : InvImage r f a₁ a₂) (h₂ : InvImage r f a₂ a₃) ↦ h h₁ h₂ + +@[deprecated "No deprecation message was provided." (since := "2024-09-13")] +theorem InvImage.irreflexive (h : Irreflexive r) : Irreflexive (InvImage r f) := + fun (a : α) (h₁ : InvImage r f a a) ↦ h (f a) h₁ + +end + +end + +/-! ### Minimal and maximal -/ + +section LE + +variable {α : Type*} [LE α] {P : α → Prop} {x y : α} + +/-- `Minimal P x` means that `x` is a minimal element satisfying `P`. -/ +def Minimal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → y ≤ x → x ≤ y + +/-- `Maximal P x` means that `x` is a maximal element satisfying `P`. -/ +def Maximal (P : α → Prop) (x : α) : Prop := P x ∧ ∀ ⦃y⦄, P y → x ≤ y → y ≤ x + +lemma Minimal.prop (h : Minimal P x) : P x := + h.1 + +lemma Maximal.prop (h : Maximal P x) : P x := + h.1 + +lemma Minimal.le_of_le (h : Minimal P x) (hy : P y) (hle : y ≤ x) : x ≤ y := + h.2 hy hle + +lemma Maximal.le_of_ge (h : Maximal P x) (hy : P y) (hge : x ≤ y) : y ≤ x := + h.2 hy hge + +end LE + +/-! ### Upper and lower sets -/ + +/-- An upper set in an order `α` is a set such that any element greater than one of its members is +also a member. Also called up-set, upward-closed set. -/ +def IsUpperSet {α : Type*} [LE α] (s : Set α) : Prop := + ∀ ⦃a b : α⦄, a ≤ b → a ∈ s → b ∈ s + +/-- A lower set in an order `α` is a set such that any element less than one of its members is also +a member. Also called down-set, downward-closed set. -/ +def IsLowerSet {α : Type*} [LE α] (s : Set α) : Prop := + ∀ ⦃a b : α⦄, b ≤ a → a ∈ s → b ∈ s + +@[inherit_doc IsUpperSet] +structure UpperSet (α : Type*) [LE α] where + /-- The carrier of an `UpperSet`. -/ + carrier : Set α + /-- The carrier of an `UpperSet` is an upper set. -/ + upper' : IsUpperSet carrier + +extend_docs UpperSet before "The type of upper sets of an order." + +@[inherit_doc IsLowerSet] +structure LowerSet (α : Type*) [LE α] where + /-- The carrier of a `LowerSet`. -/ + carrier : Set α + /-- The carrier of a `LowerSet` is a lower set. -/ + lower' : IsLowerSet carrier + +extend_docs LowerSet before "The type of lower sets of an order." diff --git a/Mathlib/Order/Interval/Set/Defs.lean b/Mathlib/Order/Interval/Set/Defs.lean index a233ef08b886c..1ca919215430a 100644 --- a/Mathlib/Order/Interval/Set/Defs.lean +++ b/Mathlib/Order/Interval/Set/Defs.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Johannes Hölzl -/ import Mathlib.Data.Set.Defs -import Mathlib.Order.Defs +import Mathlib.Order.Defs.PartialOrder /-! # Intervals diff --git a/Mathlib/Tactic/GCongr/Core.lean b/Mathlib/Tactic/GCongr/Core.lean index 90aeed7ba3815..f225f2c67bf80 100644 --- a/Mathlib/Tactic/GCongr/Core.lean +++ b/Mathlib/Tactic/GCongr/Core.lean @@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro, Heather Macbeth -/ import Lean -import Mathlib.Order.Defs -import Mathlib.Tactic.Core -import Mathlib.Tactic.GCongr.ForwardAttr import Batteries.Lean.Except import Batteries.Tactic.Exact +import Mathlib.Tactic.Core +import Mathlib.Tactic.GCongr.ForwardAttr +import Mathlib.Order.Defs.PartialOrder /-! # The `gcongr` ("generalized congruence") tactic diff --git a/Mathlib/Tactic/PushNeg.lean b/Mathlib/Tactic/PushNeg.lean index 07699624b4fc5..39b7ff1d2cd6b 100644 --- a/Mathlib/Tactic/PushNeg.lean +++ b/Mathlib/Tactic/PushNeg.lean @@ -7,7 +7,7 @@ Authors: Patrick Massot, Simon Hudon, Alice Laroche, Frédéric Dupuis, Jireh Lo import Lean.Elab.Tactic.Location import Mathlib.Data.Set.Defs import Mathlib.Logic.Basic -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.Conv /-! diff --git a/MathlibTest/push_neg.lean b/MathlibTest/push_neg.lean index 571b3a5f85cab..ccfb7e34cc4a6 100644 --- a/MathlibTest/push_neg.lean +++ b/MathlibTest/push_neg.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alice Laroche, Frédéric Dupuis, Jireh Loreaux -/ -import Mathlib.Order.Defs +import Mathlib.Order.Defs.LinearOrder import Mathlib.Tactic.PushNeg private axiom test_sorry : ∀ {α}, α diff --git a/scripts/noshake.json b/scripts/noshake.json index b3c884f2347c8..2ca0558661bde 100644 --- a/scripts/noshake.json +++ b/scripts/noshake.json @@ -250,7 +250,8 @@ "Mathlib.Tactic.IrreducibleDef": ["Mathlib.Data.Subtype"], "Mathlib.Tactic.ITauto": ["Batteries.Tactic.Init", "Mathlib.Logic.Basic"], "Mathlib.Tactic.Group": ["Mathlib.Algebra.Group.Commutator"], - "Mathlib.Tactic.GCongr.Core": ["Mathlib.Order.Defs"], + "Mathlib.Tactic.GCongr.Core": + ["Mathlib.Order.Defs", "Mathlib.Order.Defs.PartialOrder"], "Mathlib.Tactic.GCongr": ["Mathlib.Tactic.Positivity.Core"], "Mathlib.Tactic.FunProp.RefinedDiscrTree": ["Mathlib.Algebra.Group.Pi.Basic"], "Mathlib.Tactic.FunProp.Mor": ["Mathlib.Data.FunLike.Basic"], @@ -336,6 +337,7 @@ "Mathlib.Order.RelClasses": ["Batteries.WF"], "Mathlib.Order.Filter.ListTraverse": ["Mathlib.Control.Traversable.Instances"], + "Mathlib.Order.Defs.PartialOrder": ["Batteries.Tactic.Trans"], "Mathlib.Order.Defs": ["Batteries.Tactic.Trans"], "Mathlib.Order.BoundedOrder.Basic": ["Mathlib.Tactic.Finiteness.Attr"], "Mathlib.Order.BoundedOrder": ["Mathlib.Tactic.Finiteness.Attr"], @@ -349,6 +351,7 @@ "Mathlib.Logic.Function.Defs": ["Mathlib.Tactic.AdaptationNote"], "Mathlib.Logic.Function.Basic": ["Batteries.Tactic.Init"], "Mathlib.Logic.Equiv.Defs": ["Mathlib.Data.Bool.Basic"], + "Mathlib.Logic.Basic": ["Batteries.Tactic.Trans"], "Mathlib.LinearAlgebra.Matrix.Transvection": ["Mathlib.Data.Matrix.DMatrix"], "Mathlib.LinearAlgebra.DFinsupp": ["Mathlib.LinearAlgebra.Finsupp.SumProd"], "Mathlib.LinearAlgebra.Basic": ["Mathlib.Algebra.BigOperators.Group.Finset"], From c98bb5e8c10fe7c39f86614b631dcaead6afdc09 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 26 Nov 2024 22:07:36 +1100 Subject: [PATCH 183/250] chore: adaptations for nightly-2024-11-26 From 841cc90f3826f488626a68e34117d45f83cff8e5 Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Tue, 26 Nov 2024 11:45:38 +0000 Subject: [PATCH 184/250] feat(Algebra/Category): API for `Under R` (#19058) We provide API for `Under R` where `R : CommRingCat`. We also add some `simp` lemmas for the various ring categories. --- Mathlib.lean | 1 + .../ModuleCat/Differentials/Presheaf.lean | 2 +- Mathlib/Algebra/Category/Ring/Basic.lean | 39 +++- .../Algebra/Category/Ring/Constructions.lean | 11 ++ .../Algebra/Category/Ring/Under/Basic.lean | 178 ++++++++++++++++++ Mathlib/CategoryTheory/Comma/Over.lean | 20 ++ Mathlib/RingTheory/TensorProduct/Basic.lean | 9 + 7 files changed, 254 insertions(+), 6 deletions(-) create mode 100644 Mathlib/Algebra/Category/Ring/Under/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index c81570dc36eb0..fb9f9c436a030 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -152,6 +152,7 @@ import Mathlib.Algebra.Category.Ring.FilteredColimits import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits import Mathlib.Algebra.Category.Ring.LinearAlgebra +import Mathlib.Algebra.Category.Ring.Under.Basic import Mathlib.Algebra.Category.Semigrp.Basic import Mathlib.Algebra.Central.Basic import Mathlib.Algebra.Central.Defs diff --git a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean index eef6ac3891d20..5ea1401c768c2 100644 --- a/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean +++ b/Mathlib/Algebra/Category/ModuleCat/Differentials/Presheaf.lean @@ -183,7 +183,7 @@ noncomputable def relativeDifferentials' : obj X := CommRingCat.KaehlerDifferential (φ'.app X) map f := CommRingCat.KaehlerDifferential.map (φ'.naturality f) map_id _ := by ext; simp; rfl - map_comp _ _ := by ext; simp; rfl + map_comp _ _ := by ext; simp attribute [simp] relativeDifferentials'_obj diff --git a/Mathlib/Algebra/Category/Ring/Basic.lean b/Mathlib/Algebra/Category/Ring/Basic.lean index fe80c9b8c463c..61b53a227b0c6 100644 --- a/Mathlib/Algebra/Category/Ring/Basic.lean +++ b/Mathlib/Algebra/Category/Ring/Basic.lean @@ -98,6 +98,14 @@ theorem coe_of (R : Type u) [Semiring R] : (SemiRingCat.of R : Type u) = R := @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `SemiRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : SemiRingCat.{u}} + [Semiring X] [Semiring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `SemiRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `SemiRingCat.of X ⟶ SemiRingCat.of Y := X →+* Y`, -- but this one will fire. @@ -142,7 +150,7 @@ theorem ofHom_apply {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [Semiring R] [Semiring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl /-- Ring equivalence are isomorphisms in category of semirings @@ -221,7 +229,7 @@ theorem coe_of (R : Type u) [Ring R] : (RingCat.of R : Type u) = R := /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl -- Coercing the identity morphism, as a ring homomorphism, gives the identity function. @[simp] theorem coe_ringHom_id {X : RingCat} : @@ -247,6 +255,13 @@ theorem ofHom_apply' {R S : Type u} [Ring R] [Ring S] (f : R →+* S) (x : R) : @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `RingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : RingCat.{u}} [Ring X] [Ring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `RingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `RingCat.of X ⟶ RingCat.of Y := X →+* Y`, -- but this one will fire. @@ -340,7 +355,7 @@ theorem ofHom_apply {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+ /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [CommSemiring R] [CommSemiring S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl instance : Inhabited CommSemiRingCat := ⟨of PUnit⟩ @@ -376,6 +391,14 @@ theorem coe_of (R : Type u) [CommSemiring R] : (CommSemiRingCat.of R : Type u) = @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `CommSemiRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommSemiRingCat.{u}} + [CommSemiring X] [CommSemiring Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `CommSemiRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `CommSemiRingCat.of X ⟶ CommSemiRingCat.of Y := X →+* Y`, -- but this one will fire. @@ -491,8 +514,7 @@ theorem ofHom_apply {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x /-- A variant of `ofHom_apply` that makes `simpNF` happy -/ @[simp] theorem ofHom_apply' {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) (x : R) : - DFunLike.coe (α := R) (β := fun _ ↦ S) (ofHom f) x = f x := rfl - + @DFunLike.coe no_index _ R (fun _ ↦ S) _ (ofHom f) x = f x := rfl instance : Inhabited CommRingCat := ⟨of PUnit⟩ @@ -528,6 +550,13 @@ theorem coe_of (R : Type u) [CommRing R] : (CommRingCat.of R : Type u) = R := @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (RingHom.comp g f) := rfl +/-- Variant of `CommRingCat.coe_comp_of` for morphisms. -/ +@[simp] theorem coe_comp_of' {X Z : Type u} {Y : CommRingCat.{u}} [CommRing X] [CommRing Z] + (f : of X ⟶ Y) (g : Y ⟶ of Z) : + @DFunLike.coe no_index (of X ⟶ of Z) X (fun _ ↦ Z) _ (f ≫ g) = + @DFunLike.coe (X →+* Z) X (fun _ ↦ Z) _ (g.comp f) := by + rfl + -- Sometimes neither the `ext` lemma for `CommRingCat` nor for `RingHom` is applicable, -- because of incomplete unfolding of `CommRingCat.of X ⟶ CommRingCat.of Y := X →+* Y`, -- but this one will fire. diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index fea4b951aa7ee..ca21d51022790 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -5,6 +5,7 @@ Authors: Andrew Yang -/ import Mathlib.Algebra.Category.Ring.Instances import Mathlib.Algebra.Category.Ring.Limits +import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq import Mathlib.CategoryTheory.Limits.Shapes.StrictInitial import Mathlib.RingTheory.TensorProduct.Basic @@ -102,6 +103,16 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := rw [← h.map_mul, Algebra.TensorProduct.tmul_mul_tmul, mul_one, one_mul] rfl +lemma isPushout_tensorProduct (R A B : Type u) [CommRing R] [CommRing A] [CommRing B] + [Algebra R A] [Algebra R B] : + IsPushout (ofHom <| algebraMap R A) (ofHom <| algebraMap R B) + (ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) + (ofHom <| Algebra.TensorProduct.includeRight.toRingHom (R := R)) where + w := by + ext + simp + isColimit' := ⟨pushoutCoconeIsColimit R A B⟩ + end Pushout section Terminal diff --git a/Mathlib/Algebra/Category/Ring/Under/Basic.lean b/Mathlib/Algebra/Category/Ring/Under/Basic.lean new file mode 100644 index 0000000000000..a5e50b4d3c16b --- /dev/null +++ b/Mathlib/Algebra/Category/Ring/Under/Basic.lean @@ -0,0 +1,178 @@ +/- +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.Algebra.Category.Ring.Colimits +import Mathlib.Algebra.Category.Ring.Constructions +import Mathlib.CategoryTheory.Adjunction.Over + +/-! +# Under `CommRingCat` + +In this file we provide basic API for `Under R` when `R : CommRingCat`. `Under R` is +(equivalent to) the category of commutative `R`-algebras. For not necessarily commutative +algebras, use `AlgebraCat R` instead. +-/ + +noncomputable section + +universe u + +open TensorProduct CategoryTheory Limits + +variable {R S : CommRingCat.{u}} + +namespace CommRingCat + +instance : CoeSort (Under R) (Type u) where + coe A := A.right + +instance (A : Under R) : Algebra R A := RingHom.toAlgebra A.hom + +/-- Turn a morphism in `Under R` into an algebra homomorphism. -/ +def toAlgHom {A B : Under R} (f : A ⟶ B) : A →ₐ[R] B where + __ := f.right + commutes' a := by + have : (A.hom ≫ f.right) a = B.hom a := by simp + simpa only [Functor.const_obj_obj, Functor.id_obj, CommRingCat.comp_apply] using this + +@[simp] +lemma toAlgHom_id (A : Under R) : toAlgHom (𝟙 A) = AlgHom.id R A := rfl + +@[simp] +lemma toAlgHom_comp {A B C : Under R} (f : A ⟶ B) (g : B ⟶ C) : + toAlgHom (f ≫ g) = (toAlgHom g).comp (toAlgHom f) := rfl + +@[simp] +lemma toAlgHom_apply {A B : Under R} (f : A ⟶ B) (a : A) : + toAlgHom f a = f.right a := + rfl + +variable (R) in +/-- Make an object of `Under R` from an `R`-algebra. -/ +@[simps! hom, simps! (config := .lemmasOnly) right] +def mkUnder (A : Type u) [CommRing A] [Algebra R A] : Under R := + Under.mk (CommRingCat.ofHom <| algebraMap R A) + +@[ext] +lemma mkUnder_ext {A : Type u} [CommRing A] [Algebra R A] {B : Under R} + {f g : mkUnder R A ⟶ B} (h : ∀ a : A, f.right a = g.right a) : + f = g := by + ext x + exact h x + +end CommRingCat + +namespace AlgHom + +/-- Make a morphism in `Under R` from an algebra map. -/ +def toUnder {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f : A →ₐ[R] B) : CommRingCat.mkUnder R A ⟶ CommRingCat.mkUnder R B := + Under.homMk f.toRingHom <| by + ext a + exact f.commutes' a + +@[simp] +lemma toUnder_right {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A →ₐ[R] B) (a : A) : + f.toUnder.right a = f a := + rfl + +@[simp] +lemma toUnder_comp {A B C : Type u} [CommRing A] [CommRing B] [CommRing C] + [Algebra R A] [Algebra R B] [Algebra R C] (f : A →ₐ[R] B) (g : B →ₐ[R] C) : + (g.comp f).toUnder = f.toUnder ≫ g.toUnder := + rfl + +end AlgHom + +namespace AlgEquiv + +/-- Make an isomorphism in `Under R` from an algebra isomorphism. -/ +def toUnder {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] + (f : A ≃ₐ[R] B) : + CommRingCat.mkUnder R A ≅ CommRingCat.mkUnder R B where + hom := f.toAlgHom.toUnder + inv := f.symm.toAlgHom.toUnder + hom_inv_id := by + ext (a : (CommRingCat.mkUnder R A).right) + simp + inv_hom_id := by + ext a + simp + +@[simp] +lemma toUnder_hom_right_apply {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A ≃ₐ[R] B) (a : A) : + f.toUnder.hom.right a = f a := rfl + +@[simp] +lemma toUnder_inv_right_apply {A B : Type u} [CommRing A] [CommRing B] [Algebra R A] + [Algebra R B] (f : A ≃ₐ[R] B) (b : B) : + f.toUnder.inv.right b = f.symm b := rfl + +@[simp] +lemma toUnder_trans {A B C : Type u} [CommRing A] [CommRing B] [CommRing C] + [Algebra R A] [Algebra R B] [Algebra R C] (f : A ≃ₐ[R] B) (g : B ≃ₐ[R] C) : + (f.trans g).toUnder = f.toUnder ≪≫ g.toUnder := + rfl + +end AlgEquiv + +namespace CommRingCat + +variable [Algebra R S] + +variable (R S) in +/-- The base change functor `A ↦ S ⊗[R] A`. -/ +@[simps! map_right] +def tensorProd : Under R ⥤ Under S where + obj A := mkUnder S (S ⊗[R] A) + map f := Algebra.TensorProduct.map (AlgHom.id S S) (toAlgHom f) |>.toUnder + map_comp {X Y Z} f g := by simp [Algebra.TensorProduct.map_id_comp] + +variable (S) in +/-- The natural isomorphism `S ⊗[R] A ≅ pushout A.hom (algebraMap R S)` in `Under S`. -/ +def tensorProdObjIsoPushoutObj (A : Under R) : + mkUnder S (S ⊗[R] A) ≅ (Under.pushout (algebraMap R S)).obj A := + Under.isoMk (CommRingCat.isPushout_tensorProduct R S A).flip.isoPushout <| by + simp only [Functor.const_obj_obj, Under.pushout_obj, Functor.id_obj, Under.mk_right, + mkUnder_hom, AlgHom.toRingHom_eq_coe, IsPushout.inr_isoPushout_hom, Under.mk_hom] + rfl + +@[reassoc (attr := simp)] +lemma pushout_inl_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : + pushout.inl A.hom (algebraMap R S) ≫ (tensorProdObjIsoPushoutObj S A).inv.right = + (CommRingCat.ofHom <| Algebra.TensorProduct.includeRight.toRingHom) := by + simp [tensorProdObjIsoPushoutObj] + +@[reassoc (attr := simp)] +lemma pushout_inr_tensorProdObjIsoPushoutObj_inv_right (A : Under R) : + pushout.inr A.hom (algebraMap R S) ≫ + (tensorProdObjIsoPushoutObj S A).inv.right = + (CommRingCat.ofHom <| Algebra.TensorProduct.includeLeftRingHom) := by + simp [tensorProdObjIsoPushoutObj] + +variable (R S) in +/-- `A ↦ S ⊗[R] A` is naturally isomorphic to `A ↦ pushout A.hom (algebraMap R S)`. -/ +def tensorProdIsoPushout : tensorProd R S ≅ Under.pushout (algebraMap R S) := + NatIso.ofComponents (fun A ↦ tensorProdObjIsoPushoutObj S A) <| by + intro A B f + dsimp + rw [← cancel_epi (tensorProdObjIsoPushoutObj S A).inv] + ext : 1 + apply pushout.hom_ext + · rw [← cancel_mono (tensorProdObjIsoPushoutObj S B).inv.right] + ext x + simp [mkUnder_right] + · rw [← cancel_mono (tensorProdObjIsoPushoutObj S B).inv.right] + ext (x : S) + simp [mkUnder_right] + +@[simp] +lemma tensorProdIsoPushout_app (A : Under R) : + (tensorProdIsoPushout R S).app A = tensorProdObjIsoPushoutObj S A := + rfl + +end CommRingCat diff --git a/Mathlib/CategoryTheory/Comma/Over.lean b/Mathlib/CategoryTheory/Comma/Over.lean index 305664eb69cad..9f9609c579450 100644 --- a/Mathlib/CategoryTheory/Comma/Over.lean +++ b/Mathlib/CategoryTheory/Comma/Over.lean @@ -110,6 +110,16 @@ def isoMk {f g : Over X} (hl : f.left ≅ g.left) (hw : hl.hom ≫ g.hom = f.hom -- Porting note: simp solves this; simpNF still sees them after `-simp` (?) attribute [-simp, nolint simpNF] isoMk_hom_right_down_down isoMk_inv_right_down_down +@[reassoc (attr := simp)] +lemma hom_left_inv_left {f g : Over X} (e : f ≅ g) : + e.hom.left ≫ e.inv.left = 𝟙 f.left := by + simp [← Over.comp_left] + +@[reassoc (attr := simp)] +lemma inv_left_hom_left {f g : Over X} (e : f ≅ g) : + e.inv.left ≫ e.hom.left = 𝟙 g.left := by + simp [← Over.comp_left] + section variable (X) @@ -478,6 +488,16 @@ theorem isoMk_inv_right {f g : Under X} (hr : f.right ≅ g.right) (hw : f.hom (isoMk hr hw).inv.right = hr.inv := rfl +@[reassoc (attr := simp)] +lemma hom_right_inv_right {f g : Under X} (e : f ≅ g) : + e.hom.right ≫ e.inv.right = 𝟙 f.right := by + simp [← Under.comp_right] + +@[reassoc (attr := simp)] +lemma inv_right_hom_right {f g : Under X} (e : f ≅ g) : + e.inv.right ≫ e.hom.right = 𝟙 g.right := by + simp [← Under.comp_right] + section variable (X) diff --git a/Mathlib/RingTheory/TensorProduct/Basic.lean b/Mathlib/RingTheory/TensorProduct/Basic.lean index 58b8d3ab83f95..857c189049c66 100644 --- a/Mathlib/RingTheory/TensorProduct/Basic.lean +++ b/Mathlib/RingTheory/TensorProduct/Basic.lean @@ -902,6 +902,15 @@ theorem map_comp [Algebra S C] [IsScalarTower R S C] map (f₂.comp f₁) (g₂.comp g₁) = (map f₂ g₂).comp (map f₁ g₁) := ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) +lemma map_id_comp (g₂ : E →ₐ[R] F) (g₁ : D →ₐ[R] E) : + map (AlgHom.id S A) (g₂.comp g₁) = (map (AlgHom.id S A) g₂).comp (map (AlgHom.id S A) g₁) := + ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) + +lemma map_comp_id [Algebra S C] [IsScalarTower R S C] + (f₂ : B →ₐ[S] C) (f₁ : A →ₐ[S] B) : + map (f₂.comp f₁) (AlgHom.id R E) = (map f₂ (AlgHom.id R E)).comp (map f₁ (AlgHom.id R E)) := + ext (AlgHom.ext fun _ => rfl) (AlgHom.ext fun _ => rfl) + @[simp] theorem map_comp_includeLeft (f : A →ₐ[S] B) (g : C →ₐ[R] D) : (map f g).comp includeLeft = includeLeft.comp f := From 00cad59469294e95e01861e18a0026801fad3c4e Mon Sep 17 00:00:00 2001 From: grunweg Date: Tue, 26 Nov 2024 12:08:54 +0000 Subject: [PATCH 185/250] chore(100.yml): various clean-ups (#19503) Standardise how multiple authors are listed, by - removing a 'mathlib (author names)' identifier: it is apparent when results are in mathlib - always using `and` to separate multiple authors of the same formalisation (the previous state used a mix of `and` and `,`, yielding inconsistent results on the webpage). Keep using `,` to separate authors of different formalisations. --- docs/100.yaml | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/docs/100.yaml b/docs/100.yaml index 550cf7682b392..92f8c616c82ff 100644 --- a/docs/100.yaml +++ b/docs/100.yaml @@ -35,7 +35,7 @@ 9: title : The Area of a Circle decl : Theorems100.area_disc - author : James Arthur, Benjamin Davidson, and Andrew Souther + author : James Arthur and Benjamin Davidson and Andrew Souther 10: title : Euler’s Generalization of Fermat’s Little Theorem decl : Nat.ModEq.pow_totient @@ -51,7 +51,7 @@ 14: title : Euler’s Summation of 1 + (1/2)^2 + (1/3)^2 + …. decl : hasSum_zeta_two - author : Marc Masdeu, David Loeffler + author : Marc Masdeu and David Loeffler 15: title : Fundamental Theorem of Integral Calculus decls : @@ -113,7 +113,7 @@ title : Feuerbach’s Theorem 30: title : The Ballot Problem - author : Bhavik Mehta, Kexing Ying + author : Bhavik Mehta and Kexing Ying decl : Ballot.ballot_problem 31: title : Ramsey’s Theorem @@ -127,7 +127,7 @@ 34: title : Divergence of the Harmonic Series decl : Real.tendsto_sum_range_one_div_nat_succ_atTop - author : Anatole Dedecker, Yury Kudryashov + author : Anatole Dedecker and Yury Kudryashov 35: title : Taylor’s Theorem decls : @@ -157,12 +157,12 @@ 40: title : Minkowski’s Fundamental Theorem decl : MeasureTheory.exists_ne_zero_mem_lattice_of_measure_mul_two_pow_lt_measure - author : Alex J. Best, Yaël Dillies + author : Alex J. Best and Yaël Dillies 41: title : Puiseux’s Theorem 42: title : Sum of the Reciprocals of the Triangular Numbers - author : Jalex Stark, Yury Kudryashov + author : Jalex Stark and Yury Kudryashov decl : Theorems100.inverse_triangle_sum 43: title : The Isoperimetric Theorem @@ -172,7 +172,7 @@ author : Chris Hughes 45: title : The Partition Theorem - author : Bhavik Mehta, Aaron Anderson + author : Bhavik Mehta and Aaron Anderson decl : Theorems100.partition_theorem 46: title : The Solution of the General Quartic Equation @@ -310,7 +310,7 @@ decls : - sum_range_pow - sum_Ico_pow - author : mathlib (Moritz Firsching, Fabian Kruse, Ashvni Narayanan) + author : Moritz Firsching and Fabian Kruse and Ashvni Narayanan 78: title : The Cauchy-Schwarz Inequality decls : @@ -320,7 +320,7 @@ 79: title : The Intermediate Value Theorem decl : intermediate_value_Icc - author : mathlib (Rob Lewis and Chris Hughes) + author : Rob Lewis and Chris Hughes 80: title : The Fundamental Theorem of Arithmetic decls : @@ -329,14 +329,14 @@ - EuclideanDomain.to_principal_ideal_domain - UniqueFactorizationMonoid - UniqueFactorizationMonoid.factors_unique - author : mathlib (Chris Hughes) + author : Chris Hughes note : "it also has a generalized version, by showing that every Euclidean domain is a unique factorization domain, and showing that the integers form a Euclidean domain." 81: title : Divergence of the Prime Reciprocal Series decls : - Theorems100.Real.tendsto_sum_one_div_prime_atTop - not_summable_one_div_on_primes - author : Manuel Candales (archive), Michael Stoll (Mathlib) + author : Manuel Candales (archive), Michael Stoll (mathlib) 82: title : Dissection of Cubes (J.E. Littlewood’s ‘elegant’ proof) decl : Theorems100.«82».cannot_cube_a_cube @@ -344,7 +344,7 @@ 83: title : The Friendship Theorem decl : Theorems100.friendship_theorem - author : Aaron Anderson, Jalex Stark, Kyle Miller + author : Aaron Anderson and Jalex Stark and Kyle Miller 84: title : Morley’s Theorem 85: @@ -374,7 +374,7 @@ title : Stirling’s Formula decls : - Stirling.tendsto_stirlingSeq_sqrt_pi - author : mathlib (Moritz Firsching, Fabian Kruse, Nikolas Kuhn, Heather Macbeth) + author : Moritz Firsching and Fabian Kruse and Nikolas Kuhn and Heather Macbeth 91: title : The Triangle Inequality decl : norm_add_le @@ -412,7 +412,7 @@ 98: title : Bertrand’s Postulate decl : Nat.bertrand - author : Bolton Bailey, Patrick Stevens + author : Bolton Bailey and Patrick Stevens 99: title : Buffon Needle Problem links : From 7727e723796108bd9eda38dfd44c634f63ee4393 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Tue, 26 Nov 2024 12:25:44 +0000 Subject: [PATCH 186/250] chore(*): golf away from `ENNReal.toReal_le_toReal` (#19441) In most cases, use `ENNReal.toReal_mono`, since it has 1 less `ne_top` assumption. In one case, it allows to drop an assumption. --- .../Fourier/RiemannLebesgueLemma.lean | 12 ++--- Mathlib/Data/ENNReal/Real.lean | 10 ++-- .../ConditionalExpectation/CondexpL1.lean | 9 ++-- .../ConditionalExpectation/CondexpL2.lean | 3 +- .../Function/ConditionalExpectation/Real.lean | 5 +- .../Function/LpSeminorm/CompareExp.lean | 4 +- Mathlib/MeasureTheory/Function/LpSpace.lean | 10 ++-- .../Function/UniformIntegrable.lean | 8 +--- Mathlib/MeasureTheory/Integral/Bochner.lean | 7 ++- .../MeasureTheory/Integral/SetIntegral.lean | 7 ++- Mathlib/MeasureTheory/Integral/SetToL1.lean | 12 ++--- .../Measure/LevyProkhorovMetric.lean | 19 +++----- .../MeasureTheory/Measure/Portmanteau.lean | 5 +- Mathlib/Probability/BorelCantelli.lean | 6 +-- .../Kernel/Disintegration/CondCDF.lean | 6 +-- .../Kernel/Disintegration/Density.lean | 47 +++++-------------- Mathlib/Probability/Kernel/Integral.lean | 5 +- Mathlib/Probability/Moments.lean | 4 +- Mathlib/Topology/Instances/ENNReal.lean | 12 ++--- Mathlib/Topology/MetricSpace/Bounded.lean | 3 +- .../MetricSpace/HausdorffDistance.lean | 15 ++---- 21 files changed, 71 insertions(+), 138 deletions(-) diff --git a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean index 1531edb4f4be0..08b29d45a4d25 100644 --- a/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean +++ b/Mathlib/Analysis/Fourier/RiemannLebesgueLemma.lean @@ -170,20 +170,16 @@ theorem tendsto_integral_exp_inner_smul_cocompact_of_continuous_compact_support have bdA2 := norm_setIntegral_le_of_norm_le_const (hB_vol.trans_lt ENNReal.coe_lt_top) bdA ?_ swap · apply Continuous.aestronglyMeasurable - exact - continuous_norm.comp <| - Continuous.sub hf1 <| Continuous.comp hf1 <| continuous_id'.add continuous_const + exact continuous_norm.comp <| hf1.sub <| hf1.comp <| continuous_id'.add continuous_const have : ‖_‖ = ∫ v : V in A, ‖f v - f (v + i w)‖ := Real.norm_of_nonneg (setIntegral_nonneg mA fun x _ => norm_nonneg _) rw [this] at bdA2 refine bdA2.trans_lt ?_ rw [div_mul_eq_mul_div, div_lt_iff₀ (NNReal.coe_pos.mpr hB_pos), mul_comm (2 : ℝ), mul_assoc, mul_lt_mul_left hε] - rw [← ENNReal.toReal_le_toReal] at hB_vol - · refine hB_vol.trans_lt ?_ - rw [(by rfl : (↑B : ENNReal).toReal = ↑B), two_mul] - exact lt_add_of_pos_left _ hB_pos - exacts [(hB_vol.trans_lt ENNReal.coe_lt_top).ne, ENNReal.coe_lt_top.ne] + refine (ENNReal.toReal_mono ENNReal.coe_ne_top hB_vol).trans_lt ?_ + rw [ENNReal.coe_toReal, two_mul] + exact lt_add_of_pos_left _ hB_pos variable (f) diff --git a/Mathlib/Data/ENNReal/Real.lean b/Mathlib/Data/ENNReal/Real.lean index 101673394b1c5..879ef89de544e 100644 --- a/Mathlib/Data/ENNReal/Real.lean +++ b/Mathlib/Data/ENNReal/Real.lean @@ -130,13 +130,13 @@ theorem toNNReal_lt_of_lt_coe (h : a < p) : a.toNNReal < p := theorem toReal_max (hr : a ≠ ∞) (hp : b ≠ ∞) : ENNReal.toReal (max a b) = max (ENNReal.toReal a) (ENNReal.toReal b) := (le_total a b).elim - (fun h => by simp only [h, (ENNReal.toReal_le_toReal hr hp).2 h, max_eq_right]) fun h => by - simp only [h, (ENNReal.toReal_le_toReal hp hr).2 h, max_eq_left] + (fun h => by simp only [h, ENNReal.toReal_mono hp h, max_eq_right]) fun h => by + simp only [h, ENNReal.toReal_mono hr h, max_eq_left] theorem toReal_min {a b : ℝ≥0∞} (hr : a ≠ ∞) (hp : b ≠ ∞) : ENNReal.toReal (min a b) = min (ENNReal.toReal a) (ENNReal.toReal b) := - (le_total a b).elim (fun h => by simp only [h, (ENNReal.toReal_le_toReal hr hp).2 h, min_eq_left]) - fun h => by simp only [h, (ENNReal.toReal_le_toReal hp hr).2 h, min_eq_right] + (le_total a b).elim (fun h => by simp only [h, ENNReal.toReal_mono hp h, min_eq_left]) + fun h => by simp only [h, ENNReal.toReal_mono hr h, min_eq_right] theorem toReal_sup {a b : ℝ≥0∞} : a ≠ ∞ → b ≠ ∞ → (a ⊔ b).toReal = a.toReal ⊔ b.toReal := toReal_max @@ -428,7 +428,7 @@ protected theorem trichotomy₂ {p q : ℝ≥0∞} (hpq : p ≤ q) : repeat' right have hq' : 0 < q := lt_of_lt_of_le hp hpq have hp' : p < ∞ := lt_of_le_of_lt hpq hq - simp [ENNReal.toReal_le_toReal hp'.ne hq.ne, ENNReal.toReal_pos_iff, hpq, hp, hp', hq', hq] + simp [ENNReal.toReal_mono hq.ne hpq, ENNReal.toReal_pos_iff, hp, hp', hq', hq] protected theorem dichotomy (p : ℝ≥0∞) [Fact (1 ≤ p)] : p = ∞ ∨ 1 ≤ p.toReal := haveI : p = ⊤ ∨ 0 < p.toReal ∧ 1 ≤ p.toReal := by diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean index cdd2dddd698c5..b90f7b167be62 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL1.lean @@ -119,10 +119,8 @@ theorem condexpIndL1Fin_smul' [NormedSpace ℝ F] [SMulCommClass ℝ 𝕜 F] (hs theorem norm_condexpIndL1Fin_le (hs : MeasurableSet s) (hμs : μ s ≠ ∞) (x : G) : ‖condexpIndL1Fin hm hs hμs x‖ ≤ (μ s).toReal * ‖x‖ := by - have : 0 ≤ ∫ a : α, ‖condexpIndL1Fin hm hs hμs x a‖ ∂μ := by positivity - rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, ← - ENNReal.toReal_ofReal this, - ENNReal.toReal_le_toReal ENNReal.ofReal_ne_top (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), + rw [L1.norm_eq_integral_norm, ← ENNReal.toReal_ofReal (norm_nonneg x), ← ENNReal.toReal_mul, + ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.mul_ne_top hμs ENNReal.ofReal_ne_top), ofReal_integral_norm_eq_lintegral_nnnorm] swap; · rw [← memℒp_one_iff_integrable]; exact Lp.memℒp _ have h_eq : @@ -140,8 +138,7 @@ theorem condexpIndL1Fin_disjoint_union (hs : MeasurableSet s) (ht : MeasurableSe (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne x = condexpIndL1Fin hm hs hμs x + condexpIndL1Fin hm ht hμt x := by ext1 - have hμst := - ((measure_union_le s t).trans_lt (lt_top_iff_ne_top.mpr (ENNReal.add_ne_top.mpr ⟨hμs, hμt⟩))).ne + have hμst := measure_union_ne_top hμs hμt refine (condexpIndL1Fin_ae_eq_condexpIndSMul hm (hs.union ht) hμst x).trans ?_ refine EventuallyEq.trans ?_ (Lp.coeFn_add _ _).symm have hs_eq := condexpIndL1Fin_ae_eq_condexpIndSMul hm hs hμs x diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean index 54f8441a10130..5593cd5e6e6b1 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/CondexpL2.lean @@ -106,8 +106,7 @@ alias snorm_condexpL2_le := eLpNorm_condexpL2_le theorem norm_condexpL2_coe_le (hm : m ≤ m0) (f : α →₂[μ] E) : ‖(condexpL2 E 𝕜 hm f : α →₂[μ] E)‖ ≤ ‖f‖ := by rw [Lp.norm_def, Lp.norm_def, ← lpMeas_coe] - refine (ENNReal.toReal_le_toReal ?_ (Lp.eLpNorm_ne_top _)).mpr (eLpNorm_condexpL2_le hm f) - exact Lp.eLpNorm_ne_top _ + exact ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) (eLpNorm_condexpL2_le hm f) theorem inner_condexpL2_left_eq_right (hm : m ≤ m0) {f g : α →₂[μ] E} : ⟪(condexpL2 E 𝕜 hm f : α →₂[μ] E), g⟫₂ = ⟪f, (condexpL2 E 𝕜 hm g : α →₂[μ] E)⟫₂ := diff --git a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean index 44d7dd73837dd..19641592ce80f 100644 --- a/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean +++ b/Mathlib/MeasureTheory/Function/ConditionalExpectation/Real.lean @@ -99,11 +99,10 @@ theorem integral_abs_condexp_le (f : α → ℝ) : ∫ x, |(μ[f|m]) x| ∂μ mul_zero] positivity rw [integral_eq_lintegral_of_nonneg_ae, integral_eq_lintegral_of_nonneg_ae] - · rw [ENNReal.toReal_le_toReal] <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] + · apply ENNReal.toReal_mono <;> simp_rw [← Real.norm_eq_abs, ofReal_norm_eq_coe_nnnorm] + · exact hfint.2.ne · rw [← eLpNorm_one_eq_lintegral_nnnorm, ← eLpNorm_one_eq_lintegral_nnnorm] exact eLpNorm_one_condexp_le_eLpNorm _ - · exact integrable_condexp.2.ne - · exact hfint.2.ne · filter_upwards with x using abs_nonneg _ · simp_rw [← Real.norm_eq_abs] exact hfint.1.norm diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean index e7723bfa145de..796c3240dc8c4 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/CompareExp.lean @@ -84,7 +84,7 @@ theorem eLpNorm_le_eLpNorm_mul_rpow_measure_univ {p q : ℝ≥0∞} (hpq : p ≤ have hp_lt_top : p < ∞ := hpq.trans_lt (lt_top_iff_ne_top.mpr hq_top) have hp_pos : 0 < p.toReal := ENNReal.toReal_pos hp0_lt.ne' hp_lt_top.ne rw [eLpNorm_eq_eLpNorm' hp0_lt.ne.symm hp_lt_top.ne, eLpNorm_eq_eLpNorm' hq0_lt.ne.symm hq_top] - have hpq_real : p.toReal ≤ q.toReal := by rwa [ENNReal.toReal_le_toReal hp_lt_top.ne hq_top] + have hpq_real : p.toReal ≤ q.toReal := ENNReal.toReal_mono hq_top hpq exact eLpNorm'_le_eLpNorm'_mul_rpow_measure_univ hp_pos hpq_real hf @[deprecated (since := "2024-07-27")] @@ -155,7 +155,7 @@ theorem Memℒp.memℒp_of_exponent_le {p q : ℝ≥0∞} [IsFiniteMeasure μ] { have hp_eq_zero : p = 0 := le_antisymm (by rwa [hq_eq_zero] at hpq) (zero_le _) rw [hp_eq_zero, ENNReal.zero_toReal] at hp_pos exact (lt_irrefl _) hp_pos - have hpq_real : p.toReal ≤ q.toReal := by rwa [ENNReal.toReal_le_toReal hp_top hq_top] + have hpq_real : p.toReal ≤ q.toReal := ENNReal.toReal_mono hq_top hpq rw [eLpNorm_eq_eLpNorm' hp0 hp_top] rw [eLpNorm_eq_eLpNorm' hq0 hq_top] at hfq_lt_top exact eLpNorm'_lt_top_of_eLpNorm'_lt_top_of_exponent_le hfq_m hfq_lt_top hp_pos.le hpq_real diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 582ddfc879f14..0d14ecf974a22 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -352,8 +352,8 @@ theorem norm_le_mul_norm_of_ae_le_mul {c : ℝ} {f : Lp E p μ} {g : Lp F p μ} theorem norm_le_norm_of_ae_le {f : Lp E p μ} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖ ≤ ‖g x‖) : ‖f‖ ≤ ‖g‖ := by - rw [norm_def, norm_def, ENNReal.toReal_le_toReal (eLpNorm_ne_top _) (eLpNorm_ne_top _)] - exact eLpNorm_mono_ae h + rw [norm_def, norm_def] + exact ENNReal.toReal_mono (eLpNorm_ne_top _) (eLpNorm_mono_ae h) theorem mem_Lp_of_nnnorm_ae_le_mul {c : ℝ≥0} {f : α →ₘ[μ] E} {g : Lp F p μ} (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ c * ‖g x‖₊) : f ∈ Lp E p μ := @@ -1619,9 +1619,9 @@ theorem ae_tendsto_of_cauchy_eLpNorm [CompleteSpace E] {f : ℕ → α → E} refine cauchySeq_of_le_tendsto_0 (fun n => (B n).toReal) ?_ ?_ · intro n m N hnN hmN specialize hx N n m hnN hmN - rw [_root_.dist_eq_norm, ← ENNReal.toReal_ofReal (norm_nonneg _), - ENNReal.toReal_le_toReal ENNReal.ofReal_ne_top (ENNReal.ne_top_of_tsum_ne_top hB N)] - rw [← ofReal_norm_eq_coe_nnnorm] at hx + rw [_root_.dist_eq_norm, + ← ENNReal.ofReal_le_iff_le_toReal (ENNReal.ne_top_of_tsum_ne_top hB N), + ofReal_norm_eq_coe_nnnorm] exact hx.le · rw [← ENNReal.zero_toReal] exact diff --git a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean index 1f933577f78c9..c7c8a68fdf2d0 100644 --- a/Mathlib/MeasureTheory/Function/UniformIntegrable.lean +++ b/Mathlib/MeasureTheory/Function/UniformIntegrable.lean @@ -517,12 +517,8 @@ theorem tendsto_Lp_finite_of_tendsto_ae_of_meas [IsFiniteMeasure μ] (hp : 1 ≤ exact ⟨0, fun n _ => by simp [h]⟩ by_cases hμ : μ = 0 · exact ⟨0, fun n _ => by simp [hμ]⟩ - have hε' : 0 < ε.toReal / 3 := - div_pos (ENNReal.toReal_pos (gt_iff_lt.1 hε).ne.symm h.ne) (by norm_num) - have hdivp : 0 ≤ 1 / p.toReal := by - refine one_div_nonneg.2 ?_ - rw [← ENNReal.zero_toReal, ENNReal.toReal_le_toReal ENNReal.zero_ne_top hp'] - exact le_trans (zero_le _) hp + have hε' : 0 < ε.toReal / 3 := div_pos (ENNReal.toReal_pos hε.ne' h.ne) (by norm_num) + have hdivp : 0 ≤ 1 / p.toReal := by positivity have hpow : 0 < measureUnivNNReal μ ^ (1 / p.toReal) := Real.rpow_pos_of_pos (measureUnivNNReal_pos hμ) _ obtain ⟨δ₁, hδ₁, heLpNorm₁⟩ := hui hε' diff --git a/Mathlib/MeasureTheory/Integral/Bochner.lean b/Mathlib/MeasureTheory/Integral/Bochner.lean index 2cf72496737ad..a3a28a6bca1d4 100644 --- a/Mathlib/MeasureTheory/Integral/Bochner.lean +++ b/Mathlib/MeasureTheory/Integral/Bochner.lean @@ -1352,10 +1352,9 @@ theorem integral_mono_measure {f : α → ℝ} {ν} (hle : μ ≤ ν) (hf : 0 (hfi : Integrable f ν) : ∫ a, f a ∂μ ≤ ∫ a, f a ∂ν := by have hfi' : Integrable f μ := hfi.mono_measure hle have hf' : 0 ≤ᵐ[μ] f := hle.absolutelyContinuous hf - rw [integral_eq_lintegral_of_nonneg_ae hf' hfi'.1, integral_eq_lintegral_of_nonneg_ae hf hfi.1, - ENNReal.toReal_le_toReal] - exacts [lintegral_mono' hle le_rfl, ((hasFiniteIntegral_iff_ofReal hf').1 hfi'.2).ne, - ((hasFiniteIntegral_iff_ofReal hf).1 hfi.2).ne] + rw [integral_eq_lintegral_of_nonneg_ae hf' hfi'.1, integral_eq_lintegral_of_nonneg_ae hf hfi.1] + refine ENNReal.toReal_mono ?_ (lintegral_mono' hle le_rfl) + exact ((hasFiniteIntegral_iff_ofReal hf).1 hfi.2).ne theorem norm_integral_le_integral_norm (f : α → G) : ‖∫ a, f a ∂μ‖ ≤ ∫ a, ‖f a‖ ∂μ := by have le_ae : ∀ᵐ a ∂μ, 0 ≤ ‖f a‖ := Eventually.of_forall fun a => norm_nonneg _ diff --git a/Mathlib/MeasureTheory/Integral/SetIntegral.lean b/Mathlib/MeasureTheory/Integral/SetIntegral.lean index 24dbbbb85ffe7..4636afb0369c5 100644 --- a/Mathlib/MeasureTheory/Integral/SetIntegral.lean +++ b/Mathlib/MeasureTheory/Integral/SetIntegral.lean @@ -1019,10 +1019,9 @@ theorem Lp_toLp_restrict_smul (c : 𝕜) (f : Lp F p μ) (s : Set X) : `(Lp.memℒp f).restrict s).toLp f`. This map is non-expansive. -/ theorem norm_Lp_toLp_restrict_le (s : Set X) (f : Lp E p μ) : ‖((Lp.memℒp f).restrict s).toLp f‖ ≤ ‖f‖ := by - rw [Lp.norm_def, Lp.norm_def, ENNReal.toReal_le_toReal (Lp.eLpNorm_ne_top _) - (Lp.eLpNorm_ne_top _)] - apply (le_of_eq _).trans (eLpNorm_mono_measure _ (Measure.restrict_le_self (s := s))) - exact eLpNorm_congr_ae (Memℒp.coeFn_toLp _) + rw [Lp.norm_def, Lp.norm_def, eLpNorm_congr_ae (Memℒp.coeFn_toLp _)] + refine ENNReal.toReal_mono (Lp.eLpNorm_ne_top _) ?_ + exact eLpNorm_mono_measure _ Measure.restrict_le_self variable (X F 𝕜) in /-- Continuous linear map sending a function of `Lp F p μ` to the same function in diff --git a/Mathlib/MeasureTheory/Integral/SetToL1.lean b/Mathlib/MeasureTheory/Integral/SetToL1.lean index 456069af75266..2a0d2d70bd546 100644 --- a/Mathlib/MeasureTheory/Integral/SetToL1.lean +++ b/Mathlib/MeasureTheory/Integral/SetToL1.lean @@ -1422,20 +1422,14 @@ theorem continuous_L1_toL1 {μ' : Measure α} (c' : ℝ≥0∞) (hc' : c' ≠ rw [this] have h_eLpNorm_ne_top : eLpNorm (⇑g - ⇑f) 1 μ ≠ ∞ := by rw [← eLpNorm_congr_ae (Lp.coeFn_sub _ _)]; exact Lp.eLpNorm_ne_top _ - have h_eLpNorm_ne_top' : eLpNorm (⇑g - ⇑f) 1 μ' ≠ ∞ := by - refine ((eLpNorm_mono_measure _ hμ'_le).trans_lt ?_).ne - rw [eLpNorm_smul_measure_of_ne_zero hc'0, smul_eq_mul] - refine ENNReal.mul_lt_top ?_ h_eLpNorm_ne_top.lt_top - simp [hc'.lt_top, hc'0] calc (eLpNorm (⇑g - ⇑f) 1 μ').toReal ≤ (c' * eLpNorm (⇑g - ⇑f) 1 μ).toReal := by - rw [toReal_le_toReal h_eLpNorm_ne_top' (ENNReal.mul_ne_top hc' h_eLpNorm_ne_top)] - refine (eLpNorm_mono_measure (⇑g - ⇑f) hμ'_le).trans ?_ + refine toReal_mono (ENNReal.mul_ne_top hc' h_eLpNorm_ne_top) ?_ + refine (eLpNorm_mono_measure (⇑g - ⇑f) hμ'_le).trans_eq ?_ rw [eLpNorm_smul_measure_of_ne_zero hc'0, smul_eq_mul] simp _ = c'.toReal * (eLpNorm (⇑g - ⇑f) 1 μ).toReal := toReal_mul - _ ≤ c'.toReal * (ε / 2 / c'.toReal) := - (mul_le_mul le_rfl hfg.le toReal_nonneg toReal_nonneg) + _ ≤ c'.toReal * (ε / 2 / c'.toReal) := by gcongr _ = ε / 2 := by refine mul_div_cancel₀ (ε / 2) ?_; rw [Ne, toReal_eq_zero_iff]; simp [hc', hc'0] _ < ε := half_lt_self hε_pos diff --git a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean index 02ed4b6805a38..3fb7d9b953875 100644 --- a/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean +++ b/Mathlib/MeasureTheory/Measure/LevyProkhorovMetric.lean @@ -180,14 +180,10 @@ lemma levyProkhorovDist_comm (μ ν : Measure Ω) : lemma levyProkhorovDist_triangle [OpensMeasurableSpace Ω] (μ ν κ : Measure Ω) [IsFiniteMeasure μ] [IsFiniteMeasure ν] [IsFiniteMeasure κ] : levyProkhorovDist μ κ ≤ levyProkhorovDist μ ν + levyProkhorovDist ν κ := by - have dμκ_finite := (levyProkhorovEDist_lt_top μ κ).ne have dμν_finite := (levyProkhorovEDist_lt_top μ ν).ne have dνκ_finite := (levyProkhorovEDist_lt_top ν κ).ne - convert (ENNReal.toReal_le_toReal (a := levyProkhorovEDist μ κ) - (b := levyProkhorovEDist μ ν + levyProkhorovEDist ν κ) - _ _).mpr <| levyProkhorovEDist_triangle μ ν κ + convert ENNReal.toReal_mono ?_ <| levyProkhorovEDist_triangle μ ν κ · simp only [levyProkhorovDist, ENNReal.toReal_add dμν_finite dνκ_finite] - · exact dμκ_finite · exact ENNReal.add_ne_top.mpr ⟨dμν_finite, dνκ_finite⟩ /-- A type synonym, to be used for `Measure α`, `FiniteMeasure α`, or `ProbabilityMeasure α`, @@ -366,9 +362,8 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me have key : (fun (t : ℝ) ↦ ENNReal.toReal (μ {a | t ≤ f a})) ≤ (fun (t : ℝ) ↦ ENNReal.toReal (ν (thickening ε {a | t ≤ f a})) + ε) := by intro t - convert (ENNReal.toReal_le_toReal (measure_ne_top _ _) ?_).mpr - <| left_measure_le_of_levyProkhorovEDist_lt hμν (B := {a | t ≤ f a}) - (f.continuous.measurable measurableSet_Ici) + convert ENNReal.toReal_mono ?_ <| left_measure_le_of_levyProkhorovEDist_lt hμν + (B := {a | t ≤ f a}) (f.continuous.measurable measurableSet_Ici) · rw [ENNReal.toReal_add (measure_ne_top ν _) ofReal_ne_top, ENNReal.toReal_ofReal ε_pos.le] · exact ENNReal.add_ne_top.mpr ⟨measure_ne_top ν _, ofReal_ne_top⟩ have intble₁ : IntegrableOn (fun t ↦ ENNReal.toReal (μ {a | t ≤ f a})) (Ioc 0 ‖f‖) := by @@ -377,8 +372,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me exact fun _ _ hst ↦ measure_mono (fun _ h ↦ hst.trans h) · apply Eventually.of_forall <| fun t ↦ ?_ simp only [Real.norm_eq_abs, abs_toReal] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)).mpr - <| measure_mono (subset_univ _) + exact ENNReal.toReal_mono (measure_ne_top _ _) <| measure_mono (subset_univ _) have intble₂ : IntegrableOn (fun t ↦ ENNReal.toReal (ν (thickening ε {a | t ≤ f a}))) (Ioc 0 ‖f‖) := by apply Measure.integrableOn_of_bounded (M := ENNReal.toReal (ν univ)) measure_Ioc_lt_top.ne @@ -386,8 +380,7 @@ lemma BoundedContinuousFunction.integral_le_of_levyProkhorovEDist_lt (μ ν : Me exact fun _ _ hst ↦ measure_mono <| thickening_subset_of_subset ε (fun _ h ↦ hst.trans h) · apply Eventually.of_forall <| fun t ↦ ?_ simp only [Real.norm_eq_abs, abs_toReal] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)).mpr - <| measure_mono (subset_univ _) + exact ENNReal.toReal_mono (measure_ne_top _ _) <| measure_mono (subset_univ _) apply le_trans (setIntegral_mono (s := Ioc 0 ‖f‖) ?_ ?_ key) · rw [integral_add] · apply add_le_add_left @@ -415,7 +408,7 @@ lemma tendsto_integral_meas_thickening_le (f : Ω →ᵇ ℝ) · apply Eventually.of_forall (fun i ↦ ?_) apply Eventually.of_forall (fun t ↦ ?_) simp only [Real.norm_eq_abs, NNReal.abs_eq, Pi.one_apply] - exact (ENNReal.toReal_le_toReal (measure_ne_top _ _) one_ne_top).mpr prob_le_one + exact ENNReal.toReal_mono one_ne_top prob_le_one · have aux : IsFiniteMeasure (volume.restrict A) := ⟨by simp [lt_top_iff_ne_top, A_finmeas]⟩ apply integrable_const · apply Eventually.of_forall (fun t ↦ ?_) diff --git a/Mathlib/MeasureTheory/Measure/Portmanteau.lean b/Mathlib/MeasureTheory/Measure/Portmanteau.lean index 72fe3305828be..6cb6af9e29fd8 100644 --- a/Mathlib/MeasureTheory/Measure/Portmanteau.lean +++ b/Mathlib/MeasureTheory/Measure/Portmanteau.lean @@ -481,7 +481,7 @@ lemma lintegral_le_liminf_lintegral_of_forall_isOpen_measure_le_liminf_measure measure_mono (fun ω hω ↦ lt_of_le_of_lt hst hω))) lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure - {μ : Measure Ω} [IsProbabilityMeasure μ] {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] + {μ : Measure Ω} {μs : ℕ → Measure Ω} [∀ i, IsProbabilityMeasure (μs i)] {f : Ω →ᵇ ℝ} (f_nn : 0 ≤ f) (h_opens : ∀ G, IsOpen G → μ G ≤ atTop.liminf (fun i ↦ μs i G)) : ∫ x, (f x) ∂μ ≤ atTop.liminf (fun i ↦ ∫ x, (f x) ∂ (μs i)) := by @@ -489,7 +489,7 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure f.continuous f_nn h_opens rw [@integral_eq_lintegral_of_nonneg_ae Ω _ μ f (Eventually.of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] - convert (ENNReal.toReal_le_toReal ?_ ?_).mpr same + convert ENNReal.toReal_mono ?_ same · simp only [fun i ↦ @integral_eq_lintegral_of_nonneg_ae Ω _ (μs i) f (Eventually.of_forall f_nn) f.continuous.measurable.aestronglyMeasurable] let g := BoundedContinuousFunction.comp _ Real.lipschitzWith_toNNReal f @@ -497,7 +497,6 @@ lemma integral_le_liminf_integral_of_forall_isOpen_measure_le_liminf_measure simpa only [coe_nnreal_ennreal_nndist, measure_univ, mul_one, ge_iff_le] using BoundedContinuousFunction.lintegral_le_edist_mul (μ := μs i) g apply ENNReal.liminf_toReal_eq ENNReal.coe_ne_top (Eventually.of_forall bound) - · exact (f.lintegral_of_real_lt_top μ).ne · apply ne_of_lt have obs := fun (i : ℕ) ↦ @BoundedContinuousFunction.lintegral_nnnorm_le Ω _ _ (μs i) ℝ _ f simp only [measure_univ, mul_one] at obs diff --git a/Mathlib/Probability/BorelCantelli.lean b/Mathlib/Probability/BorelCantelli.lean index be4a2299f904e..c94f855abd159 100644 --- a/Mathlib/Probability/BorelCantelli.lean +++ b/Mathlib/Probability/BorelCantelli.lean @@ -92,9 +92,9 @@ theorem measure_limsup_eq_one {s : ℕ → Set Ω} (hsm : ∀ n, MeasurableSet ( · refine ⟨n, ?_⟩ rw [ENNReal.toReal_sum] exact fun _ _ => measure_ne_top _ _ - · rw [not_lt, ← ENNReal.toReal_le_toReal (ENNReal.sum_ne_top.2 _) ENNReal.coe_ne_top] - · exact hB.trans (by simp) - · exact fun _ _ => measure_ne_top _ _ + · rwa [not_lt, ENNReal.ofNNReal_toNNReal, ENNReal.le_ofReal_iff_toReal_le] + · simp + · exact le_trans (by positivity) hB end BorelCantelli diff --git a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean index b52e446ac654c..9d6ed97e8708e 100644 --- a/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean +++ b/Mathlib/Probability/Kernel/Disintegration/CondCDF.lean @@ -208,10 +208,8 @@ lemma isRatCondKernelCDFAux_preCDF (ρ : Measure (α × ℝ)) [IsFiniteMeasure (Kernel.const Unit ρ) (Kernel.const Unit ρ.fst) where measurable := measurable_preCDF'.comp measurable_snd mono' a r r' hrr' := by - filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ] with a h1 h2 - have h_ne_top : ∀ r, preCDF ρ r a ≠ ∞ := fun r ↦ ((h2 r).trans_lt ENNReal.one_lt_top).ne - rw [ENNReal.toReal_le_toReal (h_ne_top _) (h_ne_top _)] - exact h1 hrr' + filter_upwards [monotone_preCDF ρ, preCDF_le_one ρ] with a h₁ h₂ + exact ENNReal.toReal_mono ((h₂ _).trans_lt ENNReal.one_lt_top).ne (h₁ hrr') nonneg' _ q := by simp le_one' a q := by simp only [Kernel.const_apply, forall_const] diff --git a/Mathlib/Probability/Kernel/Disintegration/Density.lean b/Mathlib/Probability/Kernel/Disintegration/Density.lean index a3451dffcbecf..94af1027b716b 100644 --- a/Mathlib/Probability/Kernel/Disintegration/Density.lean +++ b/Mathlib/Probability/Kernel/Disintegration/Density.lean @@ -309,18 +309,11 @@ lemma densityProcess_mono_set (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ {s s' : Set β} (h : s ⊆ s') : densityProcess κ ν n a x s ≤ densityProcess κ ν n a x s' := by unfold densityProcess - by_cases h0 : ν a (countablePartitionSet n x) = 0 - · rw [h0, ENNReal.toReal_div, ENNReal.toReal_div] - simp - have h_ne_top : ∀ s, - κ a (countablePartitionSet n x ×ˢ s) / ν a (countablePartitionSet n x) ≠ ⊤ := by - intro s - rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact meas_countablePartitionSet_le_of_fst_le hκν n a x s - rw [ENNReal.toReal_le_toReal (h_ne_top s) (h_ne_top s')] - gcongr + obtain h₀ | h₀ := eq_or_ne (ν a (countablePartitionSet n x)) 0 + · simp [h₀] + · gcongr + simp only [ne_eq, ENNReal.div_eq_top, h₀, and_false, false_or, not_and, not_not] + exact eq_top_mono (meas_countablePartitionSet_le_of_fst_le hκν n a x s') lemma densityProcess_mono_kernel_left {κ' : Kernel α (γ × β)} (hκκ' : κ ≤ κ') (hκ'ν : fst κ' ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : @@ -331,16 +324,10 @@ lemma densityProcess_mono_kernel_left {κ' : Kernel α (γ × β)} (hκκ' : κ simp have h_le : κ' a (countablePartitionSet n x ×ˢ s) ≤ ν a (countablePartitionSet n x) := meas_countablePartitionSet_le_of_fst_le hκ'ν n a x s - rw [ENNReal.toReal_le_toReal] - · gcongr - exact hκκ' _ _ - · rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact (hκκ' _ _).trans h_le - · rw [ne_eq, ENNReal.div_eq_top] - simp only [ne_eq, h0, and_false, false_or, not_and, not_not] + gcongr + · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top + · apply hκκ' lemma densityProcess_antitone_kernel_right {ν' : Kernel α γ} (hνν' : ν ≤ ν') (hκν : fst κ ≤ ν) (n : ℕ) (a : α) (x : γ) (s : Set β) : @@ -350,16 +337,10 @@ lemma densityProcess_antitone_kernel_right {ν' : Kernel α γ} meas_countablePartitionSet_le_of_fst_le hκν n a x s by_cases h0 : ν a (countablePartitionSet n x) = 0 · simp [le_antisymm (h_le.trans h0.le) zero_le', h0] - have h0' : ν' a (countablePartitionSet n x) ≠ 0 := - fun h ↦ h0 (le_antisymm ((hνν' _ _).trans h.le) zero_le') - rw [ENNReal.toReal_le_toReal] - · gcongr - exact hνν' _ _ - · simp only [ne_eq, ENNReal.div_eq_top, h0', and_false, false_or, not_and, not_not] - refine fun h_top ↦ eq_top_mono ?_ h_top - exact h_le.trans (hνν' _ _) + gcongr · simp only [ne_eq, ENNReal.div_eq_top, h0, and_false, false_or, not_and, not_not] exact fun h_top ↦ eq_top_mono h_le h_top + · apply hνν' @[simp] lemma densityProcess_empty (κ : Kernel α (γ × β)) (ν : Kernel α γ) (n : ℕ) (a : α) (x : γ) : @@ -384,10 +365,7 @@ lemma tendsto_densityProcess_atTop_empty_of_antitone (κ : Kernel α (γ × β)) have : Tendsto (fun m ↦ κ a (countablePartitionSet n x ×ˢ seq m)) atTop (𝓝 ((κ a) (⋂ n_1, countablePartitionSet n x ×ˢ seq n_1))) := by apply tendsto_measure_iInter_atTop - · intro - -- TODO: why doesn't `measurability` work without this hint? - apply MeasurableSet.nullMeasurableSet - measurability + · measurability · exact fun _ _ h ↦ prod_mono_right <| hseq h · exact ⟨0, measure_ne_top _ _⟩ simpa only [← prod_iInter, hseq_iInter] using this @@ -603,8 +581,7 @@ lemma setIntegral_density (hκν : fst κ ≤ ν) [IsFiniteKernel ν] ENNReal.toReal_sub_of_le (measure_mono (by intro x; simp)) (measure_ne_top _ _)] rw [eq_tsub_iff_add_eq_of_le, add_comm] · exact h - · rw [ENNReal.toReal_le_toReal (measure_ne_top _ _) (measure_ne_top _ _)] - exact measure_mono (by intro x; simp) + · gcongr <;> simp · intro f hf_disj hf h_eq rw [integral_iUnion _ hf_disj (integrable_density hκν _ hs).integrableOn] · simp_rw [h_eq] diff --git a/Mathlib/Probability/Kernel/Integral.lean b/Mathlib/Probability/Kernel/Integral.lean index e97ef25407719..d586ae6a66158 100644 --- a/Mathlib/Probability/Kernel/Integral.lean +++ b/Mathlib/Probability/Kernel/Integral.lean @@ -26,9 +26,8 @@ lemma IsFiniteKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] refine Integrable.mono' (integrable_const (IsFiniteKernel.bound κ).toReal) ((κ.measurable_coe hs).ennreal_toReal.aestronglyMeasurable) (ae_of_all μ fun x ↦ ?_) - rw [Real.norm_eq_abs, abs_of_nonneg ENNReal.toReal_nonneg, - ENNReal.toReal_le_toReal (measure_ne_top _ _) (IsFiniteKernel.bound_ne_top _)] - exact Kernel.measure_le_bound _ _ _ + rw [Real.norm_eq_abs, abs_of_nonneg ENNReal.toReal_nonneg] + exact ENNReal.toReal_mono (IsFiniteKernel.bound_ne_top _) (Kernel.measure_le_bound _ _ _) lemma IsMarkovKernel.integrable (μ : Measure α) [IsFiniteMeasure μ] (κ : Kernel α β) [IsMarkovKernel κ] {s : Set β} (hs : MeasurableSet s) : diff --git a/Mathlib/Probability/Moments.lean b/Mathlib/Probability/Moments.lean index 1433dc8119e43..5cd7261ecfd42 100644 --- a/Mathlib/Probability/Moments.lean +++ b/Mathlib/Probability/Moments.lean @@ -289,8 +289,8 @@ theorem measure_ge_le_exp_mul_mgf [IsFiniteMeasure μ] (ε : ℝ) (ht : 0 ≤ t) rcases ht.eq_or_lt with ht_zero_eq | ht_pos · rw [ht_zero_eq.symm] simp only [neg_zero, zero_mul, exp_zero, mgf_zero', one_mul] - rw [ENNReal.toReal_le_toReal (measure_ne_top μ _) (measure_ne_top μ _)] - exact measure_mono (Set.subset_univ _) + gcongr + exacts [measure_ne_top _ _, Set.subset_univ _] calc (μ {ω | ε ≤ X ω}).toReal = (μ {ω | exp (t * ε) ≤ exp (t * X ω)}).toReal := by congr with ω diff --git a/Mathlib/Topology/Instances/ENNReal.lean b/Mathlib/Topology/Instances/ENNReal.lean index 029ada7f6a78a..418685fd8c5d4 100644 --- a/Mathlib/Topology/Instances/ENNReal.lean +++ b/Mathlib/Topology/Instances/ENNReal.lean @@ -1318,19 +1318,17 @@ lemma truncateToReal_eq_toReal {t x : ℝ≥0∞} (t_ne_top : t ≠ ∞) (x_le : lemma truncateToReal_le {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) {x : ℝ≥0∞} : truncateToReal t x ≤ t.toReal := by rw [truncateToReal] - apply (toReal_le_toReal _ t_ne_top).mpr (min_le_left t x) - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] + gcongr + exacts [t_ne_top, min_le_left t x] lemma truncateToReal_nonneg {t x : ℝ≥0∞} : 0 ≤ truncateToReal t x := toReal_nonneg /-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is monotone when `t ≠ ∞`. -/ lemma monotone_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Monotone (truncateToReal t) := by intro x y x_le_y - have obs_x : min t x ≠ ∞ := by - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] - have obs_y : min t y ≠ ∞ := by - simp_all only [ne_eq, min_eq_top, false_and, not_false_eq_true] - exact (ENNReal.toReal_le_toReal obs_x obs_y).mpr (min_le_min_left t x_le_y) + simp only [truncateToReal] + gcongr + exact ne_top_of_le_ne_top t_ne_top (min_le_left _ _) /-- The truncated cast `ENNReal.truncateToReal t : ℝ≥0∞ → ℝ` is continuous when `t ≠ ∞`. -/ lemma continuous_truncateToReal {t : ℝ≥0∞} (t_ne_top : t ≠ ∞) : Continuous (truncateToReal t) := by diff --git a/Mathlib/Topology/MetricSpace/Bounded.lean b/Mathlib/Topology/MetricSpace/Bounded.lean index ebe5246e8b3d2..ab85f93d8dea6 100644 --- a/Mathlib/Topology/MetricSpace/Bounded.lean +++ b/Mathlib/Topology/MetricSpace/Bounded.lean @@ -392,8 +392,7 @@ theorem diam_le_of_forall_dist_le_of_nonempty (hs : s.Nonempty) {C : ℝ} theorem dist_le_diam_of_mem' (h : EMetric.diam s ≠ ⊤) (hx : x ∈ s) (hy : y ∈ s) : dist x y ≤ diam s := by rw [diam, dist_edist] - rw [ENNReal.toReal_le_toReal (edist_ne_top _ _) h] - exact EMetric.edist_le_diam_of_mem hx hy + exact ENNReal.toReal_mono h <| EMetric.edist_le_diam_of_mem hx hy /-- Characterize the boundedness of a set in terms of the finiteness of its emetric.diameter. -/ theorem isBounded_iff_ediam_ne_top : IsBounded s ↔ EMetric.diam s ≠ ⊤ := diff --git a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean index 698c804dacc7e..cb1ae246d0e34 100644 --- a/Mathlib/Topology/MetricSpace/HausdorffDistance.lean +++ b/Mathlib/Topology/MetricSpace/HausdorffDistance.lean @@ -665,24 +665,15 @@ theorem hausdorffDist_empty' : hausdorffDist ∅ s = 0 := by simp [hausdorffDist in each set to the other set -/ theorem hausdorffDist_le_of_infDist {r : ℝ} (hr : 0 ≤ r) (H1 : ∀ x ∈ s, infDist x t ≤ r) (H2 : ∀ x ∈ t, infDist x s ≤ r) : hausdorffDist s t ≤ r := by - by_cases h1 : hausdorffEdist s t = ⊤ - · rwa [hausdorffDist, h1, ENNReal.top_toReal] rcases s.eq_empty_or_nonempty with hs | hs · rwa [hs, hausdorffDist_empty'] rcases t.eq_empty_or_nonempty with ht | ht · rwa [ht, hausdorffDist_empty] have : hausdorffEdist s t ≤ ENNReal.ofReal r := by apply hausdorffEdist_le_of_infEdist _ _ - · intro x hx - have I := H1 x hx - rwa [infDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal (infEdist_ne_top ht) ENNReal.ofReal_ne_top] at I - · intro x hx - have I := H2 x hx - rwa [infDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal (infEdist_ne_top hs) ENNReal.ofReal_ne_top] at I - rwa [hausdorffDist, ← ENNReal.toReal_ofReal hr, - ENNReal.toReal_le_toReal h1 ENNReal.ofReal_ne_top] + · simpa only [infDist, ← ENNReal.le_ofReal_iff_toReal_le (infEdist_ne_top ht) hr] using H1 + · simpa only [infDist, ← ENNReal.le_ofReal_iff_toReal_le (infEdist_ne_top hs) hr] using H2 + exact ENNReal.toReal_le_of_le_ofReal hr this /-- Bounding the Hausdorff distance by exhibiting, for any point in each set, another point in the other set at controlled distance -/ From 5d0ac0445585738ce54244d28a68e9abd4efffa5 Mon Sep 17 00:00:00 2001 From: Dagur Asgeirsson Date: Tue, 26 Nov 2024 12:47:44 +0000 Subject: [PATCH 187/250] feat(CategoryTheory): split up a product into a binary product (#19199) --- Mathlib.lean | 1 + .../CategoryTheory/Limits/Shapes/PiProd.lean | 69 +++++++++++++++++++ 2 files changed, 70 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean diff --git a/Mathlib.lean b/Mathlib.lean index fb9f9c436a030..3741c55a437ae 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1862,6 +1862,7 @@ import Mathlib.CategoryTheory.Limits.Shapes.Kernels import Mathlib.CategoryTheory.Limits.Shapes.Multiequalizer import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Basic import Mathlib.CategoryTheory.Limits.Shapes.NormalMono.Equalizers +import Mathlib.CategoryTheory.Limits.Shapes.PiProd import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.CategoryTheory.Limits.Shapes.Pullback.Assoc import Mathlib.CategoryTheory.Limits.Shapes.Pullback.CommSq diff --git a/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean new file mode 100644 index 0000000000000..6027232e5928d --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Shapes/PiProd.lean @@ -0,0 +1,69 @@ +/- +Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Dagur Asgeirsson +-/ +import Mathlib.CategoryTheory.Limits.Shapes.BinaryProducts +import Mathlib.CategoryTheory.Limits.Shapes.Products +/-! + +# A product as a binary product + +We write a product indexed by `I` as a binary product of the products indexed by a subset of `I` +and its complement. + +-/ + +namespace CategoryTheory.Limits + +variable {C I : Type*} [Category C] {X Y : I → C} + (f : (i : I) → X i ⟶ Y i) (P : I → Prop) [∀ i, Decidable (P i)] + [HasProduct X] [HasProduct Y] + [HasProduct (fun (i : {x : I // P x}) ↦ X i.val)] + [HasProduct (fun (i : {x : I // ¬ P x}) ↦ X i.val)] + [HasProduct (fun (i : {x : I // P x}) ↦ Y i.val)] + [HasProduct (fun (i : {x : I // ¬ P x}) ↦ Y i.val)] + +variable (X) in +/-- +The projection maps of a product to the products indexed by a subset and its complement, as a +binary fan. +-/ +noncomputable def Pi.binaryFanOfProp : BinaryFan (∏ᶜ (fun (i : {x : I // P x}) ↦ X i.val)) + (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := + BinaryFan.mk (P := ∏ᶜ X) (Pi.map' Subtype.val fun _ ↦ 𝟙 _) + (Pi.map' Subtype.val fun _ ↦ 𝟙 _) + +variable (X) in +/-- +A product indexed by `I` is a binary product of the products indexed by a subset of `I` and its +complement. +-/ +noncomputable def Pi.binaryFanOfPropIsLimit : IsLimit (Pi.binaryFanOfProp X P) := + BinaryFan.isLimitMk + (fun s ↦ Pi.lift fun b ↦ if h : P b then + s.π.app ⟨WalkingPair.left⟩ ≫ Pi.π (fun (i : {x : I // P x}) ↦ X i.val) ⟨b, h⟩ else + s.π.app ⟨WalkingPair.right⟩ ≫ Pi.π (fun (i : {x : I // ¬ P x}) ↦ X i.val) ⟨b, h⟩) + (by aesop) (by aesop) + (fun _ _ h₁ h₂ ↦ Pi.hom_ext _ _ fun b ↦ by + by_cases h : P b + · simp [← h₁, dif_pos h] + · simp [← h₂, dif_neg h]) + +lemma hasBinaryProduct_of_products : HasBinaryProduct (∏ᶜ (fun (i : {x : I // P x}) ↦ X i.val)) + (∏ᶜ (fun (i : {x : I // ¬ P x}) ↦ X i.val)) := + ⟨Pi.binaryFanOfProp X P, Pi.binaryFanOfPropIsLimit X P⟩ + +attribute [local instance] hasBinaryProduct_of_products + +lemma Pi.map_eq_prod_map : Pi.map f = + ((Pi.binaryFanOfPropIsLimit X P).conePointUniqueUpToIso (prodIsProd _ _)).hom ≫ + prod.map (Pi.map (fun (i : {x : I // P x}) ↦ f i.val)) + (Pi.map (fun (i : {x : I // ¬ P x}) ↦ f i.val)) ≫ + ((Pi.binaryFanOfPropIsLimit Y P).conePointUniqueUpToIso (prodIsProd _ _)).inv := by + rw [← Category.assoc, Iso.eq_comp_inv] + dsimp only [IsLimit.conePointUniqueUpToIso, binaryFanOfProp, prodIsProd] + apply prod.hom_ext + all_goals aesop_cat + +end CategoryTheory.Limits From b4b0314cc962a0b9d4cac132e7a1f742d7e640f3 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 26 Nov 2024 12:47:45 +0000 Subject: [PATCH 188/250] feat: more `eLpNorm` API (#19293) Make more arguments to rewriting lemmas explicit, more arguments to the other lemmas implicit. Add various lemmas. From LeanAPAP --- .../Function/LpSeminorm/Basic.lean | 96 ++++++++++++++----- Mathlib/MeasureTheory/Function/LpSpace.lean | 4 +- 2 files changed, 73 insertions(+), 27 deletions(-) diff --git a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean index 71c2449fa8236..4c9f13b1bf99b 100644 --- a/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean +++ b/Mathlib/MeasureTheory/Function/LpSeminorm/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Rémy Degenne, Sébastien Gouëzel -/ import Mathlib.Analysis.NormedSpace.IndicatorFunction +import Mathlib.Data.Fintype.Order import Mathlib.MeasureTheory.Function.EssSup import Mathlib.MeasureTheory.Function.AEEqFun import Mathlib.MeasureTheory.Function.SpecialFunctions.Basic @@ -37,7 +38,7 @@ noncomputable section open TopologicalSpace MeasureTheory Filter -open scoped NNReal ENNReal Topology +open scoped NNReal ENNReal Topology ComplexConjugate variable {α E F G : Type*} {m m0 : MeasurableSpace α} {p : ℝ≥0∞} {q : ℝ} {μ ν : Measure α} [NormedAddCommGroup E] [NormedAddCommGroup F] [NormedAddCommGroup G] @@ -304,19 +305,23 @@ end Zero section Neg @[simp] -theorem eLpNorm'_neg {f : α → F} : eLpNorm' (-f) q μ = eLpNorm' f q μ := by simp [eLpNorm'] +theorem eLpNorm'_neg (f : α → F) (q : ℝ) (μ : Measure α) : eLpNorm' (-f) q μ = eLpNorm' f q μ := by + simp [eLpNorm'] @[deprecated (since := "2024-07-27")] alias snorm'_neg := eLpNorm'_neg @[simp] -theorem eLpNorm_neg {f : α → F} : eLpNorm (-f) p μ = eLpNorm f p μ := by +theorem eLpNorm_neg (f : α → F) (p : ℝ≥0∞) (μ : Measure α) : eLpNorm (-f) p μ = eLpNorm f p μ := by by_cases h0 : p = 0 · simp [h0] by_cases h_top : p = ∞ · simp [h_top, eLpNormEssSup] simp [eLpNorm_eq_eLpNorm' h0 h_top] +lemma eLpNorm_sub_comm (f g : α → E) (p : ℝ≥0∞) (μ : Measure α) : + eLpNorm (f - g) p μ = eLpNorm (g - f) p μ := by simp [← eLpNorm_neg (f := f - g)] + @[deprecated (since := "2024-07-27")] alias snorm_neg := eLpNorm_neg @@ -446,6 +451,8 @@ theorem memℒp_const_iff {p : ℝ≥0∞} {c : E} (hp_ne_zero : p ≠ 0) (hp_ne end Const +variable {f : α → F} + lemma eLpNorm'_mono_nnnorm_ae {f : α → F} {g : α → G} (hq : 0 ≤ q) (h : ∀ᵐ x ∂μ, ‖f x‖₊ ≤ ‖g x‖₊) : eLpNorm' f q μ ≤ eLpNorm' g q μ := by simp only [eLpNorm'] @@ -828,8 +835,9 @@ private theorem eLpNorm_smul_measure_of_ne_zero_of_ne_top {p : ℝ≥0∞} (hp_n @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_zero_of_ne_top := eLpNorm_smul_measure_of_ne_zero_of_ne_top -theorem eLpNorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ≥0∞} (hc : c ≠ 0) : - eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by +/-- See `eLpNorm_smul_measure_of_ne_zero'` for a version with scalar multiplication by `ℝ≥0`. -/ +theorem eLpNorm_smul_measure_of_ne_zero {c : ℝ≥0∞} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) + (μ : Measure α) : eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] by_cases hp_top : p = ∞ @@ -839,12 +847,24 @@ theorem eLpNorm_smul_measure_of_ne_zero {p : ℝ≥0∞} {f : α → F} {c : ℝ @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_zero := eLpNorm_smul_measure_of_ne_zero -theorem eLpNorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) {f : α → F} (c : ℝ≥0∞) : +/-- See `eLpNorm_smul_measure_of_ne_zero` for a version with scalar multiplication by `ℝ≥0∞`. -/ +lemma eLpNorm_smul_measure_of_ne_zero' {c : ℝ≥0} (hc : c ≠ 0) (f : α → F) (p : ℝ≥0∞) + (μ : Measure α) : eLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • eLpNorm f p μ := + (eLpNorm_smul_measure_of_ne_zero (ENNReal.coe_ne_zero.2 hc) ..).trans (by simp; norm_cast) + +/-- See `eLpNorm_smul_measure_of_ne_top'` for a version with scalar multiplication by `ℝ≥0`. -/ +theorem eLpNorm_smul_measure_of_ne_top {p : ℝ≥0∞} (hp_ne_top : p ≠ ∞) (f : α → F) (c : ℝ≥0∞) : eLpNorm f p (c • μ) = c ^ (1 / p).toReal • eLpNorm f p μ := by by_cases hp0 : p = 0 · simp [hp0] · exact eLpNorm_smul_measure_of_ne_zero_of_ne_top hp0 hp_ne_top c +/-- See `eLpNorm_smul_measure_of_ne_top'` for a version with scalar multiplication by `ℝ≥0∞`. -/ +lemma eLpNorm_smul_measure_of_ne_top' (hp : p ≠ ∞) (c : ℝ≥0) (f : α → F) : + eLpNorm f p (c • μ) = c ^ p.toReal⁻¹ • eLpNorm f p μ := by + have : 0 ≤ p.toReal⁻¹ := by positivity + refine (eLpNorm_smul_measure_of_ne_top hp ..).trans ?_ + simp [ENNReal.smul_def, ENNReal.coe_rpow_of_nonneg, this] @[deprecated (since := "2024-07-27")] alias snorm_smul_measure_of_ne_top := eLpNorm_smul_measure_of_ne_top @@ -987,12 +1007,37 @@ theorem ae_le_eLpNormEssSup {f : α → F} : ∀ᵐ y ∂μ, ‖f y‖₊ ≤ eL @[deprecated (since := "2024-07-27")] alias ae_le_snormEssSup := ae_le_eLpNormEssSup +lemma eLpNormEssSup_lt_top_iff_isBoundedUnder : + eLpNormEssSup f μ < ⊤ ↔ IsBoundedUnder (· ≤ ·) (ae μ) fun x ↦ ‖f x‖₊ where + mp h := ⟨(eLpNormEssSup f μ).toNNReal, by + simp_rw [← ENNReal.coe_le_coe, ENNReal.coe_toNNReal h.ne]; exact ae_le_eLpNormEssSup⟩ + mpr := by rintro ⟨C, hC⟩; exact eLpNormEssSup_lt_top_of_ae_nnnorm_bound (C := C) hC + theorem meas_eLpNormEssSup_lt {f : α → F} : μ { y | eLpNormEssSup f μ < ‖f y‖₊ } = 0 := meas_essSup_lt @[deprecated (since := "2024-07-27")] alias meas_snormEssSup_lt := meas_eLpNormEssSup_lt +lemma eLpNorm_lt_top_of_finite [Finite α] [IsFiniteMeasure μ] : eLpNorm f p μ < ∞ := by + obtain rfl | hp₀ := eq_or_ne p 0 + · simp + obtain rfl | hp := eq_or_ne p ∞ + · simp only [eLpNorm_exponent_top, eLpNormEssSup_lt_top_iff_isBoundedUnder] + exact .le_of_finite + rw [eLpNorm_lt_top_iff_lintegral_rpow_nnnorm_lt_top hp₀ hp] + refine IsFiniteMeasure.lintegral_lt_top_of_bounded_to_ennreal μ ?_ + simp_rw [← ENNReal.coe_rpow_of_nonneg _ ENNReal.toReal_nonneg] + norm_cast + exact Finite.exists_le _ + +@[simp] lemma Memℒp.of_discrete [DiscreteMeasurableSpace α] [Finite α] [IsFiniteMeasure μ] : + Memℒp f p μ := + let ⟨C, hC⟩ := Finite.exists_le (‖f ·‖₊); .of_bound .of_finite C <| .of_forall hC + +@[simp] lemma eLpNorm_of_isEmpty [IsEmpty α] (f : α → E) (p : ℝ≥0∞) : eLpNorm f p μ = 0 := by + simp [Subsingleton.elim f 0] + lemma eLpNormEssSup_piecewise {s : Set α} (f g : α → E) [DecidablePred (· ∈ s)] (hs : MeasurableSet s) : eLpNormEssSup (Set.piecewise s f g) μ @@ -1196,37 +1241,33 @@ In this section we show inequalities on the norm. section BoundedSMul variable {𝕜 : Type*} [NormedRing 𝕜] [MulActionWithZero 𝕜 E] [MulActionWithZero 𝕜 F] -variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] +variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜 F] {c : 𝕜} {f : α → F} -theorem eLpNorm'_const_smul_le (c : 𝕜) (f : α → F) (hq_pos : 0 < q) : - eLpNorm' (c • f) q μ ≤ ‖c‖₊ • eLpNorm' f q μ := - eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le _ _) - hq_pos +theorem eLpNorm'_const_smul_le (hq : 0 < q) : eLpNorm' (c • f) q μ ≤ ‖c‖₊ • eLpNorm' f q μ := + eLpNorm'_le_nnreal_smul_eLpNorm'_of_ae_le_mul (Eventually.of_forall fun _ => nnnorm_smul_le ..) hq @[deprecated (since := "2024-07-27")] alias snorm'_const_smul_le := eLpNorm'_const_smul_le -theorem eLpNormEssSup_const_smul_le (c : 𝕜) (f : α → F) : - eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := +theorem eLpNormEssSup_const_smul_le : eLpNormEssSup (c • f) μ ≤ ‖c‖₊ • eLpNormEssSup f μ := eLpNormEssSup_le_nnreal_smul_eLpNormEssSup_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul_le := eLpNormEssSup_const_smul_le -theorem eLpNorm_const_smul_le (c : 𝕜) (f : α → F) : eLpNorm (c • f) p μ ≤ ‖c‖₊ • eLpNorm f p μ := +theorem eLpNorm_const_smul_le : eLpNorm (c • f) p μ ≤ ‖c‖₊ • eLpNorm f p μ := eLpNorm_le_nnreal_smul_eLpNorm_of_ae_le_mul (Eventually.of_forall fun _ => by simp [nnnorm_smul_le]) _ @[deprecated (since := "2024-07-27")] alias snorm_const_smul_le := eLpNorm_const_smul_le -theorem Memℒp.const_smul {f : α → E} (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (c • f) p μ := +theorem Memℒp.const_smul (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (c • f) p μ := ⟨AEStronglyMeasurable.const_smul hf.1 c, - (eLpNorm_const_smul_le c f).trans_lt (ENNReal.mul_lt_top ENNReal.coe_lt_top hf.2)⟩ + eLpNorm_const_smul_le.trans_lt (ENNReal.mul_lt_top ENNReal.coe_lt_top hf.2)⟩ -theorem Memℒp.const_mul {R} [NormedRing R] {f : α → R} (hf : Memℒp f p μ) (c : R) : - Memℒp (fun x => c * f x) p μ := +theorem Memℒp.const_mul {f : α → 𝕜} (hf : Memℒp f p μ) (c : 𝕜) : Memℒp (fun x => c * f x) p μ := hf.const_smul c end BoundedSMul @@ -1246,9 +1287,8 @@ theorem eLpNorm'_const_smul {f : α → F} (c : 𝕜) (hq_pos : 0 < q) : eLpNorm' (c • f) q μ = ‖c‖₊ • eLpNorm' f q μ := by obtain rfl | hc := eq_or_ne c 0 · simp [eLpNorm', hq_pos] - refine le_antisymm (eLpNorm'_const_smul_le _ _ hq_pos) ?_ - have : eLpNorm' _ q μ ≤ _ := eLpNorm'_const_smul_le c⁻¹ (c • f) hq_pos - rwa [inv_smul_smul₀ hc, nnnorm_inv, le_inv_smul_iff_of_pos (nnnorm_pos.2 hc)] at this + refine le_antisymm (eLpNorm'_const_smul_le hq_pos) ?_ + simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm'_const_smul_le (c := c⁻¹) (f := c • f) hq_pos @[deprecated (since := "2024-07-27")] alias snorm'_const_smul := eLpNorm'_const_smul @@ -1260,17 +1300,20 @@ theorem eLpNormEssSup_const_smul (c : 𝕜) (f : α → F) : @[deprecated (since := "2024-07-27")] alias snormEssSup_const_smul := eLpNormEssSup_const_smul -theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) : +theorem eLpNorm_const_smul (c : 𝕜) (f : α → F) (p : ℝ≥0∞) (μ : Measure α): eLpNorm (c • f) p μ = (‖c‖₊ : ℝ≥0∞) * eLpNorm f p μ := by obtain rfl | hc := eq_or_ne c 0 · simp - refine le_antisymm (eLpNorm_const_smul_le _ _) ?_ - have : eLpNorm _ p μ ≤ _ := eLpNorm_const_smul_le c⁻¹ (c • f) - rwa [inv_smul_smul₀ hc, nnnorm_inv, le_inv_smul_iff_of_pos (nnnorm_pos.2 hc)] at this + refine le_antisymm eLpNorm_const_smul_le ?_ + simpa [hc, le_inv_smul_iff_of_pos] using eLpNorm_const_smul_le (c := c⁻¹) (f := c • f) @[deprecated (since := "2024-07-27")] alias snorm_const_smul := eLpNorm_const_smul +lemma eLpNorm_nsmul [NormedSpace ℝ F] (n : ℕ) (f : α → F) : + eLpNorm (n • f) p μ = n * eLpNorm f p μ := by + simpa [Nat.cast_smul_eq_nsmul] using eLpNorm_const_smul (n : ℝ) f .. + end NormedSpace theorem le_eLpNorm_of_bddBelow (hp : p ≠ 0) (hp' : p ≠ ∞) {f : α → F} (C : ℝ≥0) {s : Set α} @@ -1299,6 +1342,9 @@ section RCLike variable {𝕜 : Type*} [RCLike 𝕜] {f : α → 𝕜} +@[simp] lemma eLpNorm_conj (f : α → 𝕜) (p : ℝ≥0∞) (μ : Measure α) : + eLpNorm (conj f) p μ = eLpNorm f p μ := by simp [← eLpNorm_norm] + theorem Memℒp.re (hf : Memℒp f p μ) : Memℒp (fun x => RCLike.re (f x)) p μ := by have : ∀ x, ‖RCLike.re (f x)‖ ≤ 1 * ‖f x‖ := by intro x diff --git a/Mathlib/MeasureTheory/Function/LpSpace.lean b/Mathlib/MeasureTheory/Function/LpSpace.lean index 0d14ecf974a22..ca21a90b697cf 100644 --- a/Mathlib/MeasureTheory/Function/LpSpace.lean +++ b/Mathlib/MeasureTheory/Function/LpSpace.lean @@ -426,7 +426,7 @@ variable [BoundedSMul 𝕜 E] [BoundedSMul 𝕜' E] theorem const_smul_mem_Lp (c : 𝕜) (f : Lp E p μ) : c • (f : α →ₘ[μ] E) ∈ Lp E p μ := by rw [mem_Lp_iff_eLpNorm_lt_top, eLpNorm_congr_ae (AEEqFun.coeFn_smul _ _)] - refine (eLpNorm_const_smul_le _ _).trans_lt ?_ + refine eLpNorm_const_smul_le.trans_lt ?_ rw [ENNReal.smul_def, smul_eq_mul, ENNReal.mul_lt_top_iff] exact Or.inl ⟨ENNReal.coe_lt_top, f.prop⟩ @@ -464,7 +464,7 @@ instance instBoundedSMul [Fact (1 ≤ p)] : BoundedSMul 𝕜 (Lp E p μ) := suffices (‖r • f‖₊ : ℝ≥0∞) ≤ ‖r‖₊ * ‖f‖₊ from mod_cast this rw [nnnorm_def, nnnorm_def, ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _), eLpNorm_congr_ae (coeFn_smul _ _), ENNReal.coe_toNNReal (Lp.eLpNorm_ne_top _)] - exact eLpNorm_const_smul_le r f + exact eLpNorm_const_smul_le end BoundedSMul From 5ad34c033417c6e6efd3fcd9062fa1764d240209 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 26 Nov 2024 13:02:57 +0000 Subject: [PATCH 189/250] feat(Algebra/Exact): iff lemmas (#19504) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit When we have a commutative diagram from a sequence of two maps to another, such that the left vertical map is surjective, the middle vertical map is bijective and the right vertical map is injective, then the upper row is exact iff the lower row is. This generalizes the case when all the three vertical maps are bijections. This PR introduces two new lemmas `exact_iff_of_surjective_of_bijective_of_injective` in the `AddMonoidHom` and `LinearMap` namespaces, and refactors the proofs of the all-bijections versions. This parallels the categorical lemma https://leanprover-community.github.io/mathlib4_docs/Mathlib/Algebra/Homology/ShortComplex/Exact.html#CategoryTheory.ShortComplex.exact_iff_of_epi_of_isIso_of_mono Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib/Algebra/Exact.lean | 96 ++++++++++++++++++++++++++------------ 1 file changed, 67 insertions(+), 29 deletions(-) diff --git a/Mathlib/Algebra/Exact.lean b/Mathlib/Algebra/Exact.lean index 13f90a7ac5eca..c9da18e29ffbc 100644 --- a/Mathlib/Algebra/Exact.lean +++ b/Mathlib/Algebra/Exact.lean @@ -90,6 +90,44 @@ lemma exact_of_comp_of_mem_range (h1 : g.comp f = 0) (h2 : ∀ x, g x = 0 → x ∈ range f) : Exact f g := exact_of_comp_eq_zero_of_ker_le_range h1 h2 +/-- When we have a commutative diagram from a sequence of two maps to another, +such that the left vertical map is surjective, the middle vertical map is bijective and the right +vertical map is injective, then the upper row is exact iff the lower row is. +See `ShortComplex.exact_iff_of_epi_of_isIso_of_mono` in the file +`Algebra.Homology.ShortComplex.Exact` for the categorical version of this result. -/ +lemma exact_iff_of_surjective_of_bijective_of_injective + {M₁ M₂ M₃ N₁ N₂ N₃ : Type*} [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] + (f : M₁ →+ M₂) (g : M₂ →+ M₃) (f' : N₁ →+ N₂) (g' : N₂ →+ N₃) + (τ₁ : M₁ →+ N₁) (τ₂ : M₂ →+ N₂) (τ₃ : M₃ →+ N₃) + (comm₁₂ : f'.comp τ₁ = τ₂.comp f) + (comm₂₃ : g'.comp τ₂ = τ₃.comp g) + (h₁ : Function.Surjective τ₁) (h₂ : Function.Bijective τ₂) (h₃ : Function.Injective τ₃) : + Exact f g ↔ Exact f' g' := by + replace comm₁₂ := DFunLike.congr_fun comm₁₂ + replace comm₂₃ := DFunLike.congr_fun comm₂₃ + dsimp at comm₁₂ comm₂₃ + constructor + · intro h y₂ + obtain ⟨x₂, rfl⟩ := h₂.2 y₂ + constructor + · intro hx₂ + obtain ⟨x₁, rfl⟩ := (h x₂).1 (h₃ (by simpa only [map_zero, comm₂₃] using hx₂)) + exact ⟨τ₁ x₁, by simp only [comm₁₂]⟩ + · rintro ⟨y₁, hy₁⟩ + obtain ⟨x₁, rfl⟩ := h₁ y₁ + rw [comm₂₃, (h x₂).2 _, map_zero] + exact ⟨x₁, h₂.1 (by simpa only [comm₁₂] using hy₁)⟩ + · intro h x₂ + constructor + · intro hx₂ + obtain ⟨y₁, hy₁⟩ := (h (τ₂ x₂)).1 (by simp only [comm₂₃, hx₂, map_zero]) + obtain ⟨x₁, rfl⟩ := h₁ y₁ + exact ⟨x₁, h₂.1 (by simpa only [comm₁₂] using hy₁)⟩ + · rintro ⟨x₁, rfl⟩ + apply h₃ + simp only [← comm₁₂, ← comm₂₃, h.apply_apply_eq_zero (τ₁ x₁), map_zero] + end AddMonoidHom namespace Function.Exact @@ -110,38 +148,18 @@ variable {X₁ X₂ X₃ Y₁ Y₂ Y₃ : Type*} [AddCommMonoid X₁] [AddCommMo (e₁ : X₁ ≃+ Y₁) (e₂ : X₂ ≃+ Y₂) (e₃ : X₃ ≃+ Y₃) {f₁₂ : X₁ →+ X₂} {f₂₃ : X₂ →+ X₃} {g₁₂ : Y₁ →+ Y₂} {g₂₃ : Y₂ →+ Y₃} +lemma iff_of_ladder_addEquiv (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := + (exact_iff_of_surjective_of_bijective_of_injective _ _ _ _ e₁ e₂ e₃ comm₁₂ comm₂₃ + e₁.surjective e₂.bijective e₃.injective).symm + lemma of_ladder_addEquiv_of_exact (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := by - have h₁₂ := DFunLike.congr_fun comm₁₂ - have h₂₃ := DFunLike.congr_fun comm₂₃ - dsimp at h₁₂ h₂₃ - apply of_comp_eq_zero_of_ker_in_range - · ext y₁ - obtain ⟨x₁, rfl⟩ := e₁.surjective y₁ - dsimp - rw [h₁₂, h₂₃, H.apply_apply_eq_zero, map_zero] - · intro y₂ hx₂ - obtain ⟨x₂, rfl⟩ := e₂.surjective y₂ - obtain ⟨x₁, rfl⟩ := (H x₂).1 (e₃.injective (by rw [← h₂₃, hx₂, map_zero])) - exact ⟨e₁ x₁, by rw [h₁₂]⟩ + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact f₁₂ f₂₃) : Exact g₁₂ g₂₃ := + (iff_of_ladder_addEquiv _ _ _ comm₁₂ comm₂₃).2 H lemma of_ladder_addEquiv_of_exact' (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := by - refine of_ladder_addEquiv_of_exact e₁.symm e₂.symm e₃.symm ?_ ?_ H - · ext y₁ - obtain ⟨x₁, rfl⟩ := e₁.surjective y₁ - apply e₂.injective - simpa using DFunLike.congr_fun comm₁₂.symm x₁ - · ext y₂ - obtain ⟨x₂, rfl⟩ := e₂.surjective y₂ - apply e₃.injective - simpa using DFunLike.congr_fun comm₂₃.symm x₂ - -lemma iff_of_ladder_addEquiv (comm₁₂ : g₁₂.comp e₁ = AddMonoidHom.comp e₂ f₁₂) - (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) : Exact g₁₂ g₂₃ ↔ Exact f₁₂ f₂₃ := by - constructor - · exact of_ladder_addEquiv_of_exact' e₁ e₂ e₃ comm₁₂ comm₂₃ - · exact of_ladder_addEquiv_of_exact e₁ e₂ e₃ comm₁₂ comm₂₃ + (comm₂₃ : g₂₃.comp e₂ = AddMonoidHom.comp e₃ f₂₃) (H : Exact g₁₂ g₂₃) : Exact f₁₂ f₂₃ := + (iff_of_ladder_addEquiv _ _ _ comm₁₂ comm₂₃).1 H end @@ -396,6 +414,26 @@ end Function namespace LinearMap +/-- When we have a commutative diagram from a sequence of two linear maps to another, +such that the left vertical map is surjective, the middle vertical map is bijective and the right +vertical map is injective, then the upper row is exact iff the lower row is. +See `ShortComplex.exact_iff_of_epi_of_isIso_of_mono` in the file +`Algebra.Homology.ShortComplex.Exact` for the categorical version of this result. -/ +lemma exact_iff_of_surjective_of_bijective_of_injective + {M₁ M₂ M₃ N₁ N₂ N₃ : Type*} [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] + [AddCommMonoid N₁] [AddCommMonoid N₂] [AddCommMonoid N₃] + [Module R M₁] [Module R M₂] [Module R M₃] + [Module R N₁] [Module R N₂] [Module R N₃] + (f : M₁ →ₗ[R] M₂) (g : M₂ →ₗ[R] M₃) (f' : N₁ →ₗ[R] N₂) (g' : N₂ →ₗ[R] N₃) + (τ₁ : M₁ →ₗ[R] N₁) (τ₂ : M₂ →ₗ[R] N₂) (τ₃ : M₃ →ₗ[R] N₃) + (comm₁₂ : f'.comp τ₁ = τ₂.comp f) (comm₂₃ : g'.comp τ₂ = τ₃.comp g) + (h₁ : Function.Surjective τ₁) (h₂ : Function.Bijective τ₂) (h₃ : Function.Injective τ₃) : + Function.Exact f g ↔ Function.Exact f' g' := + AddMonoidHom.exact_iff_of_surjective_of_bijective_of_injective + f.toAddMonoidHom g.toAddMonoidHom f'.toAddMonoidHom g'.toAddMonoidHom + τ₁.toAddMonoidHom τ₂.toAddMonoidHom τ₃.toAddMonoidHom + (by ext; apply DFunLike.congr_fun comm₁₂) (by ext; apply DFunLike.congr_fun comm₂₃) h₁ h₂ h₃ + lemma surjective_range_liftQ (h : range f ≤ ker g) (hg : Function.Surjective g) : Function.Surjective ((range f).liftQ g h) := by intro x₃ From 13fc0e8a77fcb3fc528f47e1e5a1b09df301bb0c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Tue, 26 Nov 2024 14:07:40 +0000 Subject: [PATCH 190/250] feat(Pointwise): monotonicity of `pow` (#19252) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit ... and `s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0` From GrowthInGroups --- Mathlib/Algebra/Algebra/Operations.lean | 2 +- .../Algebra/Group/Pointwise/Finset/Basic.lean | 134 ++++++++++++++---- .../Algebra/Group/Pointwise/Set/Basic.lean | 121 ++++++++++++---- .../Algebra/Order/Monoid/Unbundled/Pow.lean | 14 ++ Mathlib/Analysis/Convex/EGauge.lean | 3 +- Mathlib/GroupTheory/OrderOfElement.lean | 2 +- 6 files changed, 214 insertions(+), 62 deletions(-) diff --git a/Mathlib/Algebra/Algebra/Operations.lean b/Mathlib/Algebra/Algebra/Operations.lean index 960f805fd05b7..7c503761b174e 100644 --- a/Mathlib/Algebra/Algebra/Operations.lean +++ b/Mathlib/Algebra/Algebra/Operations.lean @@ -230,7 +230,7 @@ theorem pow_subset_pow {n : ℕ} : (↑M : Set A) ^ n ⊆ ↑(M ^ n : Submodule trans AddSubmonoid.pow_subset_pow (le_pow_toAddSubmonoid M) theorem pow_mem_pow {x : A} (hx : x ∈ M) (n : ℕ) : x ^ n ∈ M ^ n := - pow_subset_pow _ <| Set.pow_mem_pow hx _ + pow_subset_pow _ <| Set.pow_mem_pow hx end Module diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index ac78f4f5b38b5..876f75a97121b 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -544,14 +544,17 @@ theorem singleton_mul_singleton (a b : α) : ({a} : Finset α) * {b} = {a * b} : theorem mul_subset_mul : s₁ ⊆ s₂ → t₁ ⊆ t₂ → s₁ * t₁ ⊆ s₂ * t₂ := image₂_subset -@[to_additive] +@[to_additive (attr := gcongr)] theorem mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := image₂_subset_left -@[to_additive] +@[to_additive (attr := gcongr)] theorem mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image₂_subset_right +@[to_additive] instance : MulLeftMono (Finset α) where elim _s _t₁ _t₂ := mul_subset_mul_left +@[to_additive] instance : MulRightMono (Finset α) where elim _t _s₁ _s₂ := mul_subset_mul_right + @[to_additive] theorem mul_subset_iff : s * t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, x * y ∈ u := image₂_subset_iff @@ -947,29 +950,69 @@ protected def monoid : Monoid (Finset α) := scoped[Pointwise] attribute [instance] Finset.monoid Finset.addMonoid +-- `Finset.pow_left_monotone` doesn't exist since it would syntactically be a special case of +-- `pow_left_mono` + @[to_additive] -theorem pow_mem_pow (ha : a ∈ s) : ∀ n : ℕ, a ^ n ∈ s ^ n - | 0 => by - simp only [pow_zero, mem_one] - | n + 1 => by - simp only [pow_succ] - exact mul_mem_mul (pow_mem_pow ha n) ha +protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := + pow_right_monotone <| one_subset.2 hs + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := + Finset.pow_right_monotone hs hmn + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := + (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) + +@[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right + +@[deprecated (since := "2024-11-19")] +alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right + +@[to_additive] +lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : + s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn + +@[to_additive (attr := simp) nsmul_empty] +lemma empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] + +@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty @[to_additive] -theorem pow_subset_pow (hst : s ⊆ t) : ∀ n : ℕ, s ^ n ⊆ t ^ n - | 0 => by - simp [pow_zero] - | n + 1 => by - rw [pow_succ] - exact mul_subset_mul (pow_subset_pow hst n) hst +lemma Nonempty.pow (hs : s.Nonempty) : ∀ {n}, (s ^ n).Nonempty + | 0 => by simp + | n + 1 => by rw [pow_succ]; exact hs.pow.mul hs + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma pow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + -- TODO: The `nonempty_iff_ne_empty` would be unnecessary if `push_neg` knew how to simplify + -- `s ≠ ∅` to `s.Nonempty` when `s : Finset α`. + -- See https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/push_neg.20extensibility + · exact nonempty_iff_ne_empty.1 (nonempty_iff_ne_empty.2 hs).pow + · rw [← nonempty_iff_ne_empty] + simp + · rintro ⟨rfl, hn⟩ + exact empty_pow hn + +@[to_additive (attr := simp) nsmul_singleton] +lemma singleton_pow (a : α) : ∀ n, ({a} : Finset α) ^ n = {a ^ n} + | 0 => by simp [singleton_one] + | n + 1 => by simp [pow_succ, singleton_pow _ n] + +@[to_additive] lemma pow_mem_pow (ha : a ∈ s) : a ^ n ∈ s ^ n := by + simpa using pow_subset_pow_left (singleton_subset_iff.2 ha) + +@[to_additive] lemma one_mem_pow (hs : 1 ∈ s) : 1 ∈ s ^ n := by simpa using pow_mem_pow hs @[to_additive] -theorem pow_subset_pow_of_one_mem (hs : (1 : α) ∈ s) : m ≤ n → s ^ m ⊆ s ^ n := by - apply Nat.le_induction - · exact fun _ hn => hn - · intro n _ hmn - rw [pow_succ] - exact hmn.trans (subset_mul_left (s ^ n) hs) +lemma inter_pow_subset : (s ∩ t) ^ n ⊆ s ^ n ∩ t ^ n := by apply subset_inter <;> gcongr <;> simp @[to_additive (attr := simp, norm_cast)] theorem coe_list_prod (s : List (Finset α)) : (↑s.prod : Set α) = (s.map (↑)).prod := @@ -986,17 +1029,6 @@ theorem mem_pow {a : α} {n : ℕ} : a ∈ s ^ n ↔ ∃ f : Fin n → s, (List.ofFn fun i => ↑(f i)).prod = a := by simp [← mem_coe, coe_pow, Set.mem_pow] -@[to_additive (attr := simp) nsmul_empty] -theorem empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by - rw [← tsub_add_cancel_of_le (Nat.succ_le_of_lt <| Nat.pos_of_ne_zero hn), pow_succ', empty_mul] - -@[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty - -@[to_additive (attr := simp) nsmul_singleton] -lemma singleton_pow (a : α) : ∀ n, ({a} : Finset α) ^ n = {a ^ n} - | 0 => by simp [singleton_one] - | n + 1 => by simp [pow_succ, singleton_pow _ n] - @[to_additive] lemma card_pow_le : ∀ {n}, (s ^ n).card ≤ s.card ^ n | 0 => by simp @@ -1105,6 +1137,22 @@ lemma univ_div_univ [Fintype α] : (univ / univ : Finset α) = univ := by simp [ @[to_additive (attr := simp) zsmul_empty] lemma empty_zpow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := by cases n <;> aesop +@[to_additive] +lemma Nonempty.zpow (hs : s.Nonempty) : ∀ {n : ℤ}, (s ^ n).Nonempty + | (n : ℕ) => hs.pow + | .negSucc n => by simpa using hs.pow + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma zpow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact nonempty_iff_ne_empty.1 (nonempty_iff_ne_empty.2 hs).zpow + · rw [← nonempty_iff_ne_empty] + simp + · rintro ⟨rfl, hn⟩ + exact empty_zpow hn + @[to_additive (attr := simp) zsmul_singleton] lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Finset α) ^ n = {a ^ n} := by cases n <;> simp @@ -1129,10 +1177,18 @@ variable (f : F) {s t : Finset α} {a b : α} theorem one_mem_div_iff : (1 : α) ∈ s / t ↔ ¬Disjoint s t := by rw [← mem_coe, ← disjoint_coe, coe_div, Set.one_mem_div_iff] +@[to_additive (attr := simp)] +lemma one_mem_inv_mul_iff : (1 : α) ∈ t⁻¹ * s ↔ ¬Disjoint s t := by + aesop (add simp [not_disjoint_iff_nonempty_inter, mem_mul, mul_eq_one_iff_eq_inv, + Finset.Nonempty]) + @[to_additive] theorem not_one_mem_div_iff : (1 : α) ∉ s / t ↔ Disjoint s t := one_mem_div_iff.not_left +@[to_additive] +lemma not_one_mem_inv_mul_iff : (1 : α) ∉ t⁻¹ * s ↔ Disjoint s t := one_mem_inv_mul_iff.not_left + @[to_additive] theorem Nonempty.one_mem_div (h : s.Nonempty) : (1 : α) ∈ s / s := let ⟨a, ha⟩ := h @@ -1430,6 +1486,22 @@ theorem mul_subset_iff_right : s * t ⊆ u ↔ ∀ b ∈ t, op b • s ⊆ u := end Mul +section Monoid +variable [DecidableEq α] [DecidableEq β] [Monoid α] [Monoid β] [FunLike F α β] + +@[to_additive] +lemma image_pow_of_ne_zero [MulHomClass F α β] : + ∀ {n}, n ≠ 0 → ∀ (f : F) (s : Finset α), (s ^ n).image f = s.image f ^ n + | 1, _ => by simp + | n + 2, _ => by simp [image_mul, pow_succ _ n.succ, image_pow_of_ne_zero] + +@[to_additive] +lemma image_pow [MonoidHomClass F α β] (f : F) (s : Finset α) : ∀ n, (s ^ n).image f = s.image f ^ n + | 0 => by simp [singleton_one] + | n + 1 => image_pow_of_ne_zero n.succ_ne_zero .. + +end Monoid + section Semigroup variable [Semigroup α] [DecidableEq α] diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index 417727b33b985..a834f509b4e16 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -6,6 +6,7 @@ Authors: Johan Commelin, Floris van Doorn import Mathlib.Algebra.Group.Equiv.Basic import Mathlib.Algebra.Group.Units.Hom import Mathlib.Algebra.Opposites +import Mathlib.Algebra.Order.Monoid.Unbundled.Pow import Mathlib.Data.Set.Lattice /-! @@ -106,7 +107,7 @@ theorem one_mem_one : (1 : α) ∈ (1 : Set α) := theorem one_subset : 1 ⊆ s ↔ (1 : α) ∈ s := singleton_subset_iff -@[to_additive] +@[to_additive (attr := simp)] theorem one_nonempty : (1 : Set α).Nonempty := ⟨1, rfl⟩ @@ -347,6 +348,9 @@ theorem mul_subset_mul_left : t₁ ⊆ t₂ → s * t₁ ⊆ s * t₂ := theorem mul_subset_mul_right : s₁ ⊆ s₂ → s₁ * t ⊆ s₂ * t := image2_subset_right +@[to_additive] instance : MulLeftMono (Set α) where elim _s _t₁ _t₂ := mul_subset_mul_left +@[to_additive] instance : MulRightMono (Set α) where elim _t _s₁ _s₂ := mul_subset_mul_right + @[to_additive] theorem mul_subset_iff : s * t ⊆ u ↔ ∀ x ∈ s, ∀ y ∈ t, x * y ∈ u := image2_subset_iff @@ -1075,46 +1079,66 @@ protected noncomputable def monoid : Monoid (Set α) := scoped[Pointwise] attribute [instance] Set.monoid Set.addMonoid -@[to_additive] -theorem pow_mem_pow (ha : a ∈ s) : ∀ n : ℕ, a ^ n ∈ s ^ n - | 0 => by - simp only [pow_zero, mem_one] - | n + 1 => by - simp only [pow_succ] - exact mul_mem_mul (pow_mem_pow ha _) ha +-- `Set.pow_left_monotone` doesn't exist since it would syntactically be a special case of +-- `pow_left_mono` @[to_additive] -theorem pow_subset_pow (hst : s ⊆ t) : ∀ n : ℕ, s ^ n ⊆ t ^ n - | 0 => by - rw [pow_zero] - exact Subset.rfl - | n + 1 => by - rw [pow_succ] - exact mul_subset_mul (pow_subset_pow hst _) hst +protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := + pow_right_monotone <| one_subset.2 hs + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := + Set.pow_right_monotone hs hmn + +@[to_additive (attr := gcongr)] +lemma pow_subset_pow (hst : s ⊆ t) (ht : 1 ∈ t) (hmn : m ≤ n) : s ^ m ⊆ t ^ n := + (pow_subset_pow_left hst).trans (pow_subset_pow_right ht hmn) + +@[deprecated (since := "2024-11-19")] alias pow_subset_pow_of_one_mem := pow_subset_pow_right + +@[deprecated (since := "2024-11-19")] +alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right @[to_additive] -theorem pow_subset_pow_of_one_mem (hs : (1 : α) ∈ s) (hn : m ≤ n) : s ^ m ⊆ s ^ n := by - -- Porting note: `Nat.le_induction` didn't work as an induction principle in mathlib3, this was - -- `refine Nat.le_induction ...` - induction n, hn using Nat.le_induction with - | base => exact Subset.rfl - | succ _ _ ih => - dsimp only - rw [pow_succ'] - exact ih.trans (subset_mul_right _ hs) +lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : + s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn @[to_additive (attr := simp) nsmul_empty] -theorem empty_pow {n : ℕ} (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by - match n with - | n + 1 => rw [pow_succ', empty_mul] +lemma empty_pow (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] @[deprecated (since := "2024-10-21")] alias empty_nsmul := nsmul_empty +@[to_additive] +lemma Nonempty.pow (hs : s.Nonempty) : ∀ {n}, (s ^ n).Nonempty + | 0 => by simp + | n + 1 => by rw [pow_succ]; exact hs.pow.mul hs + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma pow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact hs.pow + · simp + · rintro ⟨rfl, hn⟩ + exact empty_pow hn + @[to_additive (attr := simp) nsmul_singleton] lemma singleton_pow (a : α) : ∀ n, ({a} : Set α) ^ n = {a ^ n} | 0 => by simp [singleton_one] | n + 1 => by simp [pow_succ, singleton_pow _ n] +@[to_additive] lemma pow_mem_pow (ha : a ∈ s) : a ^ n ∈ s ^ n := by + simpa using pow_subset_pow_left (singleton_subset_iff.2 ha) + +@[to_additive] lemma one_mem_pow (hs : 1 ∈ s) : 1 ∈ s ^ n := by simpa using pow_mem_pow hs + +@[to_additive] +lemma inter_pow_subset : (s ∩ t) ^ n ⊆ s ^ n ∩ t ^ n := by apply subset_inter <;> gcongr <;> simp + @[to_additive] theorem mul_univ_of_one_mem (hs : (1 : α) ∈ s) : s * univ = univ := eq_univ_iff_forall.2 fun _ => mem_mul.2 ⟨_, hs, _, mem_univ _, one_mul _⟩ @@ -1206,6 +1230,21 @@ lemma univ_div_univ : (univ / univ : Set α) = univ := by simp [div_eq_mul_inv] @[to_additive (attr := simp) zsmul_empty] lemma empty_zpow (hn : n ≠ 0) : (∅ : Set α) ^ n = ∅ := by cases n <;> aesop +@[to_additive] +lemma Nonempty.zpow (hs : s.Nonempty) : ∀ {n : ℤ}, (s ^ n).Nonempty + | (n : ℕ) => hs.pow + | .negSucc n => by simpa using hs.pow + +set_option push_neg.use_distrib true in +@[to_additive (attr := simp)] lemma zpow_eq_empty : s ^ n = ∅ ↔ s = ∅ ∧ n ≠ 0 := by + constructor + · contrapose! + rintro (hs | rfl) + · exact hs.zpow + · simp + · rintro ⟨rfl, hn⟩ + exact empty_zpow hn + @[to_additive (attr := simp) zsmul_singleton] lemma singleton_zpow (a : α) (n : ℤ) : ({a} : Set α) ^ n = {a ^ n} := by cases n <;> simp @@ -1231,10 +1270,17 @@ variable [Group α] {s t : Set α} {a b : α} theorem one_mem_div_iff : (1 : α) ∈ s / t ↔ ¬Disjoint s t := by simp [not_disjoint_iff_nonempty_inter, mem_div, div_eq_one, Set.Nonempty] +@[to_additive (attr := simp)] +lemma one_mem_inv_mul_iff : (1 : α) ∈ t⁻¹ * s ↔ ¬Disjoint s t := by + aesop (add simp [not_disjoint_iff_nonempty_inter, mem_mul, mul_eq_one_iff_eq_inv, Set.Nonempty]) + @[to_additive] theorem not_one_mem_div_iff : (1 : α) ∉ s / t ↔ Disjoint s t := one_mem_div_iff.not_left +@[to_additive] +lemma not_one_mem_inv_mul_iff : (1 : α) ∉ t⁻¹ * s ↔ Disjoint s t := one_mem_inv_mul_iff.not_left + alias ⟨_, _root_.Disjoint.one_not_mem_div_set⟩ := not_one_mem_div_iff attribute [to_additive] Disjoint.one_not_mem_div_set @@ -1334,6 +1380,22 @@ lemma preimage_mul (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : end Mul +section Monoid +variable [Monoid α] [Monoid β] [FunLike F α β] + +@[to_additive] +lemma image_pow_of_ne_zero [MulHomClass F α β] : + ∀ {n}, n ≠ 0 → ∀ (f : F) (s : Set α), f '' (s ^ n) = (f '' s) ^ n + | 1, _ => by simp + | n + 2, _ => by simp [image_mul, pow_succ _ n.succ, image_pow_of_ne_zero] + +@[to_additive] +lemma image_pow [MonoidHomClass F α β] (f : F) (s : Set α) : ∀ n, f '' (s ^ n) = (f '' s) ^ n + | 0 => by simp [singleton_one] + | n + 1 => image_pow_of_ne_zero n.succ_ne_zero .. + +end Monoid + section Group variable [Group α] [DivisionMonoid β] [FunLike F α β] [MonoidHomClass F α β] (m : F) {s t : Set α} @@ -1363,6 +1425,11 @@ lemma preimage_div (hm : Injective m) {s t : Set β} (hs : s ⊆ range m) (ht : end Group +variable {ι : Type*} {α : ι → Type*} [∀ i, Inv (α i)] + +@[to_additive (attr := simp)] +lemma inv_pi (s : Set ι) (t : ∀ i, Set (α i)) : (s.pi t)⁻¹ = s.pi fun i ↦ (t i)⁻¹ := by ext x; simp + end Set /-! ### Miscellaneous -/ diff --git a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean index 0a60e09d11d81..32f0ddb84013a 100644 --- a/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean +++ b/Mathlib/Algebra/Order/Monoid/Unbundled/Pow.lean @@ -119,6 +119,20 @@ theorem Right.pow_lt_one_of_lt {n : ℕ} {x : M} (hn : 0 < n) (h : x < 1) : x ^ @[deprecated (since := "2024-09-21")] alias Right.pow_neg := Right.nsmul_neg +/-- This lemma is useful in non-cancellative monoids, like sets under pointwise operations. -/ +@[to_additive +"This lemma is useful in non-cancellative monoids, like sets under pointwise operations."] +lemma pow_le_pow_mul_of_sq_le_mul [MulLeftMono M] {a b : M} (hab : a ^ 2 ≤ b * a) : + ∀ {n}, n ≠ 0 → a ^ n ≤ b ^ (n - 1) * a + | 1, _ => by simp + | n + 2, _ => by + calc + a ^ (n + 2) = a ^ (n + 1) * a := by rw [pow_succ] + _ ≤ b ^ n * a * a := mul_le_mul_right' (pow_le_pow_mul_of_sq_le_mul hab (by omega)) _ + _ = b ^ n * a ^ 2 := by rw [mul_assoc, sq] + _ ≤ b ^ n * (b * a) := mul_le_mul_left' hab _ + _ = b ^ (n + 1) * a := by rw [← mul_assoc, ← pow_succ] + end Right section CovariantLTSwap diff --git a/Mathlib/Analysis/Convex/EGauge.lean b/Mathlib/Analysis/Convex/EGauge.lean index 7ade295422705..6a61629fae9e7 100644 --- a/Mathlib/Analysis/Convex/EGauge.lean +++ b/Mathlib/Analysis/Convex/EGauge.lean @@ -111,8 +111,7 @@ lemma egauge_zero_right (hs : s.Nonempty) : egauge 𝕜 s 0 = 0 := by have : 0 ∈ (0 : 𝕜) • s := by simp [zero_smul_set hs] simpa using egauge_le_of_mem_smul this -@[simp] -lemma egauge_zero_zero : egauge 𝕜 (0 : Set E) 0 = 0 := egauge_zero_right _ ⟨0, rfl⟩ +lemma egauge_zero_zero : egauge 𝕜 (0 : Set E) 0 = 0 := by simp lemma egauge_le_one (h : x ∈ s) : egauge 𝕜 s x ≤ 1 := by rw [← one_smul 𝕜 s] at h diff --git a/Mathlib/GroupTheory/OrderOfElement.lean b/Mathlib/GroupTheory/OrderOfElement.lean index b18e30f62e781..564cc2da28dcb 100644 --- a/Mathlib/GroupTheory/OrderOfElement.lean +++ b/Mathlib/GroupTheory/OrderOfElement.lean @@ -1044,7 +1044,7 @@ def powCardSubgroup {G : Type*} [Group G] [Fintype G] (S : Set G) (hS : S.Nonemp have one_mem : (1 : G) ∈ S ^ Fintype.card G := by obtain ⟨a, ha⟩ := hS rw [← pow_card_eq_one] - exact Set.pow_mem_pow ha (Fintype.card G) + exact Set.pow_mem_pow ha subgroupOfIdempotent (S ^ Fintype.card G) ⟨1, one_mem⟩ <| by classical apply (Set.eq_of_subset_of_card_le (Set.subset_mul_left _ one_mem) (ge_of_eq _)).symm From e6b14d47d78f3846fbf3c5d50bb1eccd71abc6e5 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Tue, 26 Nov 2024 14:42:14 +0000 Subject: [PATCH 191/250] doc(GroupTheory/Commensurable): clarify some docstrings (#19127) --- Mathlib/GroupTheory/Commensurable.lean | 16 ++++++++++------ Mathlib/GroupTheory/Index.lean | 14 +++++++------- 2 files changed, 17 insertions(+), 13 deletions(-) diff --git a/Mathlib/GroupTheory/Commensurable.lean b/Mathlib/GroupTheory/Commensurable.lean index b084767be576a..e69c86c8a8949 100644 --- a/Mathlib/GroupTheory/Commensurable.lean +++ b/Mathlib/GroupTheory/Commensurable.lean @@ -8,19 +8,23 @@ import Mathlib.GroupTheory.Index /-! # Commensurability for subgroups -This file defines commensurability for subgroups of a group `G`. It then goes on to prove that -commensurability defines an equivalence relation and finally defines the commensurator of a subgroup -of `G`. +Two subgroups `H` and `K` of a group `G` are commensurable if `H ∩ K` has finite index in both `H` +and `K`. + +This file defines commensurability for subgroups of a group `G`. It goes on to prove that +commensurability defines an equivalence relation on subgroups of `G` and finally defines the +commensurator of a subgroup `H` of `G`, which is the elements `g` of `G` such that `gHg⁻¹` is +commensurable with `H`. ## Main definitions -* `Commensurable`: defines commensurability for two subgroups `H`, `K` of `G` -* `commensurator`: defines the commensurator of a subgroup `H` of `G`. +* `Commensurable H K`: the statement that the subgroups `H` and `K` of `G` are commensurable. +* `commensurator H`: the commensurator of a subgroup `H` of `G`. ## Implementation details We define the commensurator of a subgroup `H` of `G` by first defining it as a subgroup of -`(conjAct G)`, which we call commensurator' and then taking the pre-image under +`(conjAct G)`, which we call `commensurator'` and then taking the pre-image under the map `G → (conjAct G)` to obtain our commensurator as a subgroup of `G`. -/ diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index c31a41f62822f..0201d5ad30137 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -42,16 +42,16 @@ open Cardinal variable {G G' : Type*} [Group G] [Group G'] (H K L : Subgroup G) -/-- The index of a subgroup as a natural number, and returns 0 if the index is infinite. -/ -@[to_additive "The index of a subgroup as a natural number, -and returns 0 if the index is infinite."] +/-- The index of a subgroup as a natural number. Returns `0` if the index is infinite. -/ +@[to_additive "The index of an additive subgroup as a natural number. +Returns 0 if the index is infinite."] noncomputable def index : ℕ := Nat.card (G ⧸ H) -/-- The relative index of a subgroup as a natural number, - and returns 0 if the relative index is infinite. -/ -@[to_additive "The relative index of a subgroup as a natural number, -and returns 0 if the relative index is infinite."] +/-- If `H` and `K` are subgroups of a group `G`, then `relindex H K : ℕ` is the index +of `H ∩ K` in `K`. The function returns `0` if the index is infinite. -/ +@[to_additive "If `H` and `K` are subgroups of an additive group `G`, then `relindex H K : ℕ` +is the index of `H ∩ K` in `K`. The function returns `0` if the index is infinite."] noncomputable def relindex : ℕ := (H.subgroupOf K).index From 19094c4d59858e7ff3ef353faebf2e5bd009dacb Mon Sep 17 00:00:00 2001 From: Heather Macbeth <25316162+hrmacbeth@users.noreply.github.com> Date: Tue, 26 Nov 2024 17:39:51 +0000 Subject: [PATCH 192/250] feat: handle scalar multiplication in `linear_combination` (#19136) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR extends `linear_combination` to handle scalar multiplication. Examples (under sensible typeclasses, here omitted): ```lean example {a b : K} {x : V} (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by linear_combination (norm := module) h • x example {v w x y : V} (h₁ : x - y = -(v - w)) (h₂ : x + y = v + w) : x = w := by linear_combination (norm := module) (2:K)⁻¹ • (h₁ + h₂) example {a b : K} {x y z : V} (hb : 0 ≤ b) (hab : a + b = 1) (H : z ≤ a • x + b • y) (hyz : y ≤ z) : a • z ≤ a • x := by linear_combination (norm := skip) b • hyz + hab • z + H apply le_of_eq module ``` As is apparent from these examples, it is still generally necessary to invoke a nonstandard discharger (generally `module` or `match_scalars`) when using `linear_combination` with this new feature, since the default discharger (`ring`) is not usually useful on module goals. --- Mathlib/Analysis/Convex/Join.lean | 2 +- Mathlib/Analysis/Convex/Side.lean | 2 +- .../Analysis/InnerProductSpace/OfNorm.lean | 4 +- Mathlib/Tactic/LinearCombination.lean | 18 ++++-- Mathlib/Tactic/LinearCombination/Lemmas.lean | 56 +++++++++++++++++++ MathlibTest/linear_combination.lean | 53 +++++++++++++++++- MathlibTest/module.lean | 41 ++++---------- 7 files changed, 135 insertions(+), 41 deletions(-) diff --git a/Mathlib/Analysis/Convex/Join.lean b/Mathlib/Analysis/Convex/Join.lean index 058cec43d5537..76080ce818ed6 100644 --- a/Mathlib/Analysis/Convex/Join.lean +++ b/Mathlib/Analysis/Convex/Join.lean @@ -112,7 +112,7 @@ theorem convexJoin_assoc_aux (s t u : Set E) : rintro _ ⟨z, ⟨x, hx, y, hy, a₁, b₁, ha₁, hb₁, hab₁, rfl⟩, z, hz, a₂, b₂, ha₂, hb₂, hab₂, rfl⟩ obtain rfl | hb₂ := hb₂.eq_or_lt · refine ⟨x, hx, y, ⟨y, hy, z, hz, left_mem_segment 𝕜 _ _⟩, a₁, b₁, ha₁, hb₁, hab₁, ?_⟩ - linear_combination (norm := module) congr(-$hab₂ • (a₁ • x + b₁ • y)) + linear_combination (norm := module) -hab₂ • (a₁ • x + b₁ • y) refine ⟨x, hx, (a₂ * b₁ / (a₂ * b₁ + b₂)) • y + (b₂ / (a₂ * b₁ + b₂)) • z, ⟨y, hy, z, hz, _, _, by positivity, by positivity, by field_simp, rfl⟩, diff --git a/Mathlib/Analysis/Convex/Side.lean b/Mathlib/Analysis/Convex/Side.lean index 9e7135f3517d0..e623f6f31e49d 100644 --- a/Mathlib/Analysis/Convex/Side.lean +++ b/Mathlib/Analysis/Convex/Side.lean @@ -580,7 +580,7 @@ theorem wOppSide_iff_exists_wbtw {s : AffineSubspace R P} {x y : P} : · have : (r₂ / (r₁ + r₂)) • (y -ᵥ p₂ + (p₂ -ᵥ p₁) - (x -ᵥ p₁)) + (x -ᵥ p₁) = (r₂ / (r₁ + r₂)) • (p₂ -ᵥ p₁) := by rw [← neg_vsub_eq_vsub_rev p₂ y] - linear_combination (norm := match_scalars <;> field_simp) congr((r₁ + r₂)⁻¹ • $h) + linear_combination (norm := match_scalars <;> field_simp) (r₁ + r₂)⁻¹ • h rw [lineMap_apply, ← vsub_vadd x p₁, ← vsub_vadd y p₂, vsub_vadd_eq_vsub_sub, vadd_vsub_assoc, ← vadd_assoc, vadd_eq_add, this] exact s.smul_vsub_vadd_mem (r₂ / (r₁ + r₂)) hp₂ hp₁ hp₁ diff --git a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean index b62522a0e45d6..374ce175c377f 100644 --- a/Mathlib/Analysis/InnerProductSpace/OfNorm.lean +++ b/Mathlib/Analysis/InnerProductSpace/OfNorm.lean @@ -121,10 +121,10 @@ theorem inner_.conj_symm (x y : E) : conj (inner_ 𝕜 y x) = inner_ 𝕜 x y := have I_smul (v : E) : ‖(I : 𝕜) • v‖ = ‖v‖ := by rw [norm_smul, norm_I_of_ne_zero hI, one_mul] have h₁ : ‖(I : 𝕜) • y - x‖ = ‖(I : 𝕜) • x + y‖ := by convert I_smul ((I : 𝕜) • x + y) using 2 - linear_combination (norm := module) congr(-$hI' • x) + linear_combination (norm := module) -hI' • x have h₂ : ‖(I : 𝕜) • y + x‖ = ‖(I : 𝕜) • x - y‖ := by convert (I_smul ((I : 𝕜) • y + x)).symm using 2 - linear_combination (norm := module) congr(-$hI' • y) + linear_combination (norm := module) -hI' • y rw [h₁, h₂] ring diff --git a/Mathlib/Tactic/LinearCombination.lean b/Mathlib/Tactic/LinearCombination.lean index 5c891e021eae9..d5f5a2b863b2e 100644 --- a/Mathlib/Tactic/LinearCombination.lean +++ b/Mathlib/Tactic/LinearCombination.lean @@ -46,9 +46,9 @@ inductive Expanded /-- A value, equivalently a proof of `c = c`. -/ | const (c : Syntax.Term) -/-- The handling in `linear_combination` of left- and right-multiplication and of division all three -proceed according to the same logic, specified here: given a proof `p` of an (in)equality and a -constant `c`, +/-- The handling in `linear_combination` of left- and right-multiplication and scalar-multiplication +and of division all five proceed according to the same logic, specified here: given a proof `p` of +an (in)equality and a constant `c`, * if `p` is a proof of an equation, multiply/divide through by `c`; * if `p` is a proof of a non-strict inequality, run `positivity` to find a proof that `c` is nonnegative, then multiply/divide through by `c`, invoking the nonnegativity of `c` where needed; @@ -58,7 +58,7 @@ constant `c`, This generic logic takes as a parameter the object `lems`: the four lemmas corresponding to the four cases. -/ -def rescale (lems : Ineq.WithStrictness → Name) (ty : Expr) (p c : Term) : +def rescale (lems : Ineq.WithStrictness → Name) (ty : Option Expr) (p c : Term) : Ineq → TermElabM Expanded | eq => do let i := mkIdent <| lems .eq @@ -86,7 +86,8 @@ using `+`/`-`/`*`/`/` on equations and values. inequality. * `.const c` means that the input expression is not an equation but a value. -/ -partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expanded := withRef stx do +partial def expandLinearCombo (ty : Option Expr) (stx : Syntax.Term) : + TermElabM Expanded := withRef stx do match stx with | `(($e)) => expandLinearCombo ty e | `($e₁ + $e₂) => do @@ -128,6 +129,13 @@ partial def expandLinearCombo (ty : Expr) (stx : Syntax.Term) : TermElabM Expand | .const c₁, .proof rel₂ p₂ => rescale mulConstRelData ty p₂ c₁ rel₂ | .proof _ _, .proof _ _ => throwErrorAt tk "'linear_combination' supports only linear operations" + | `($e₁ •%$tk $e₂) => do + match ← expandLinearCombo none e₁, ← expandLinearCombo ty e₂ with + | .const c₁, .const c₂ => .const <$> ``($c₁ • $c₂) + | .proof rel₁ p₁, .const c₂ => rescale smulRelConstData ty p₁ c₂ rel₁ + | .const c₁, .proof rel₂ p₂ => rescale smulConstRelData none p₂ c₁ rel₂ + | .proof _ _, .proof _ _ => + throwErrorAt tk "'linear_combination' supports only linear operations" | `($e₁ /%$tk $e₂) => do match ← expandLinearCombo ty e₁, ← expandLinearCombo ty e₂ with | .const c₁, .const c₂ => .const <$> ``($c₁ / $c₂) diff --git a/Mathlib/Tactic/LinearCombination/Lemmas.lean b/Mathlib/Tactic/LinearCombination/Lemmas.lean index e22bec542ae8b..b8b2a3acda4f9 100644 --- a/Mathlib/Tactic/LinearCombination/Lemmas.lean +++ b/Mathlib/Tactic/LinearCombination/Lemmas.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.Algebra.Order.Field.Defs +import Mathlib.Algebra.Order.Module.OrderedSMul import Mathlib.Data.Ineq /-! @@ -17,6 +18,7 @@ open Lean namespace Mathlib.Tactic.LinearCombination variable {α : Type*} {a a' a₁ a₂ b b' b₁ b₂ c : α} +variable {K : Type*} {t s : K} /-! ### Addition -/ @@ -68,6 +70,42 @@ theorem mul_const_lt_weak [OrderedSemiring α] (p : b < c) {a : α} (ha : 0 ≤ a * b ≤ a * c := mul_le_mul_of_nonneg_left p.le ha +/-! ### Scalar multiplication -/ + +theorem smul_eq_const [SMul K α] (p : t = s) (c : α) : t • c = s • c := p ▸ rfl + +theorem smul_le_const [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t ≤ s) {a : α} (ha : 0 ≤ a) : + t • a ≤ s • a := + smul_le_smul_of_nonneg_right p ha + +theorem smul_lt_const [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t < s) {a : α} (ha : 0 < a) : + t • a < s • a := + smul_lt_smul_of_pos_right p ha + +theorem smul_lt_const_weak [OrderedRing K] [OrderedAddCommGroup α] [Module K α] + [OrderedSMul K α] (p : t < s) {a : α} (ha : 0 ≤ a) : + t • a ≤ s • a := + smul_le_smul_of_nonneg_right p.le ha + +theorem smul_const_eq [SMul K α] (p : b = c) (s : K) : s • b = s • c := p ▸ rfl + +theorem smul_const_le [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b ≤ c) {s : K} (hs : 0 ≤ s) : + s • b ≤ s • c := + smul_le_smul_of_nonneg_left p hs + +theorem smul_const_lt [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b < c) {s : K} (hs : 0 < s) : + s • b < s • c := + smul_lt_smul_of_pos_left p hs + +theorem smul_const_lt_weak [OrderedSemiring K] [OrderedAddCommMonoid α] [Module K α] + [OrderedSMul K α] (p : b < c) {s : K} (hs : 0 ≤ s) : + s • b ≤ s • c := + smul_le_smul_of_nonneg_left p.le hs + /-! ### Division -/ theorem div_eq_const [Div α] (p : a = b) (c : α) : a / c = b / c := p ▸ rfl @@ -176,6 +214,24 @@ def mulConstRelData : Ineq.WithStrictness → Name | .lt true => ``mul_const_lt | .lt false => ``mul_const_lt_weak +/-- Given an (in)equality, look up the lemma to left-scalar-multiply it by a constant (scalar). +If relevant, also take into account the degree of positivity which can be proved of the constant: +strict or non-strict. -/ +def smulRelConstData : Ineq.WithStrictness → Name + | .eq => ``smul_eq_const + | .le => ``smul_le_const + | .lt true => ``smul_lt_const + | .lt false => ``smul_lt_const_weak + +/-- Given an (in)equality, look up the lemma to right-scalar-multiply it by a constant (vector). +If relevant, also take into account the degree of positivity which can be proved of the constant: +strict or non-strict. -/ +def smulConstRelData : Ineq.WithStrictness → Name + | .eq => ``smul_const_eq + | .le => ``smul_const_le + | .lt true => ``smul_const_lt + | .lt false => ``smul_const_lt_weak + /-- Given an (in)equality, look up the lemma to divide it by a constant. If relevant, also take into account the degree of positivity which can be proved of the constant: strict or non-strict. -/ def divRelConstData : Ineq.WithStrictness → Name diff --git a/MathlibTest/linear_combination.lean b/MathlibTest/linear_combination.lean index 4aa730f384389..b947d8190146e 100644 --- a/MathlibTest/linear_combination.lean +++ b/MathlibTest/linear_combination.lean @@ -1,6 +1,8 @@ import Mathlib.Tactic.Abel +import Mathlib.Tactic.FieldSimp import Mathlib.Tactic.LinearCombination import Mathlib.Tactic.Linarith +import Mathlib.Tactic.Module set_option autoImplicit true @@ -123,6 +125,55 @@ example (x y z w : ℚ) (hzw : z = w) : x * z + 2 * y * z = x * w + 2 * y * w := example (x y : ℤ) (h : x = 0) : y ^ 2 * x = 0 := by linear_combination y ^ 2 * h +/-! ### Scalar multiplication -/ + +section +variable {K V : Type*} + +section +variable [AddCommGroup V] [Field K] [CharZero K] [Module K V] {a b μ ν : K} {v w x y : V} + +example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by + linear_combination (norm := module) h • x + +example (h1 : a • x + b • y = 0) (h2 : a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 := by + linear_combination (norm := module) h2 - ν • h1 + +example (h₁ : x - y = -(v - w)) (h₂ : x + y = v + w) : x = w := by + linear_combination (norm := module) (2:K)⁻¹ • h₁ + (2:K)⁻¹ • h₂ + +example (h : a + b ≠ 0) (H : a • x = b • y) : x = (b / (a + b)) • (x + y) := by + linear_combination (norm := match_scalars) (a + b)⁻¹ • H + · field_simp + ring + · ring + +end + +example [OrderedCommSemiring K] [OrderedCancelAddCommMonoid V] [Module K V] [OrderedSMul K V] + {x y r : V} (hx : x < r) (hy : y < r) {a b : K} (ha : 0 < a) (hb : 0 ≤ b) (hab : a + b = 1) : + a • x + b • y < r := by + linear_combination (norm := skip) a • hx + b • hy + hab • r + apply le_of_eq + module + +example [OrderedCommSemiring K] [OrderedCancelAddCommMonoid V] [Module K V] [OrderedSMul K V] + {x y z : V} (hyz : y ≤ z) {a b : K} (hb : 0 ≤ b) (hab : a + b = 1) (H : z ≤ a • x + b • y) : + a • z ≤ a • x := by + linear_combination (norm := skip) b • hyz + hab • z + H + apply le_of_eq + module + +example [OrderedCommRing K] [OrderedAddCommGroup V] [Module K V] [OrderedSMul K V] + {x y : V} (hx : 0 < x) (hxy : x < y) {a b c : K} (hc : 0 < c) (hac : c < a) (hab : a + b ≤ 1): + c • x + b • y < y := by + have := hx.trans hxy + linear_combination (norm := skip) hab • y + hac • y + c • hxy + apply le_of_eq + module + +end + /-! ### Tests in semirings -/ example (a _b : ℕ) (h1 : a = 3) : a = 3 := by @@ -470,7 +521,7 @@ typeclass inference is demanded by the lemmas it orchestrates. This example too (and 73 ms on a good laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of `CovariantClass`/`ContravariantClass`, and takes 206 heartbeats (10 ms on a good laptop) on the implementation at the time of joining Mathlib (November 2024). -/ -set_option maxHeartbeats 1000 in +set_option maxHeartbeats 1100 in example {a b : ℝ} (h : a < b) : 0 < b - a := by linear_combination (norm := skip) h exact test_sorry diff --git a/MathlibTest/module.lean b/MathlibTest/module.lean index 2736fdc32d143..8c96ddc566272 100644 --- a/MathlibTest/module.lean +++ b/MathlibTest/module.lean @@ -156,12 +156,8 @@ example (h : a = b) : a • x = b • x := by match_scalars linear_combination h -/- `linear_combination` does not currently handle `•`. The following mimics what should eventually -be performed by a `linear_combination` call, with exact syntax TBD -- maybe -`linear_combination (norm := module) h • x` or `module_combination h • x`. -/ example (h : a = b) : a • x = b • x := by - apply eq_of_eq (congr($h • x):) - module + linear_combination (norm := module) h • x example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by match_scalars @@ -169,25 +165,17 @@ example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b · ring example (h : a ^ 2 + b ^ 2 = 1) : a • (a • x - b • y) + (b • a • y + b • b • x) = x := by - -- `linear_combination (norm := module) h • x` - apply eq_of_eq (congr($h • x):) - module + linear_combination (norm := module) h • x example (h1 : a • x + b • y = 0) (h2 : a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 ∧ (μ - ν) • b • y = 0 := by constructor - · -- `linear_combination (norm := module) h2 - ν • h1` - apply eq_of_eq (congr($h2 - ν • $h1):) - module - · -- `linear_combination (norm := module) μ • h1 + h2` - apply eq_of_eq (congr(μ • $h1 - $h2):) - module + · linear_combination (norm := module) h2 - ν • h1 + · linear_combination (norm := module) μ • h1 - h2 example (h1 : 0 • z + a • x + b • y = 0) (h2 : 0 • ρ • z + a • μ • x + b • ν • y = 0) : (μ - ν) • a • x = 0 := by - -- `linear_combination (norm := module) h2 - ν • h1` - apply eq_of_eq (congr($h2 - ν • $h1):) - module + linear_combination (norm := module) h2 - ν • h1 example (h1 : a • x + b • y + c • z = 0) @@ -196,15 +184,9 @@ example (μ - ν) • (μ - ρ) • a • x = 0 ∧ (μ - ν) • (ν - ρ) • b • y = 0 ∧ (μ - ρ) • (ν - ρ) • c • z = 0 := by refine ⟨?_, ?_, ?_⟩ - · -- `linear_combination (norm := module) h3 - (ν + ρ) • h2 + ν • ρ • h1` - apply eq_of_eq (congr($h3 - (ν + ρ) • $h2 + ν • ρ • $h1):) - module - · -- `linear_combination (norm := module) - h3 + (μ + ρ) • h2 - μ • ρ • h1` - apply eq_of_eq (congr(- $h3 + (μ + ρ) • $h2 - μ • ρ • $h1):) - module - · -- `linear_combination (norm := module) h3 - (μ + ν) • h2 + μ • ν • h1` - apply eq_of_eq (congr($h3 - (μ + ν) • $h2 + μ • ν • $h1):) - module + · linear_combination (norm := module) h3 - (ν + ρ) • h2 + ν • ρ • h1 + · linear_combination (norm := module) - h3 + (μ + ρ) • h2 - μ • ρ • h1 + · linear_combination (norm := module) h3 - (μ + ν) • h2 + μ • ν • h1 /-- error: ring failed, ring expressions not equal @@ -253,9 +235,7 @@ variable [Field K] [CharZero K] [Module K V] example : (2:K)⁻¹ • x + (3:K)⁻¹ • x + (6:K)⁻¹ • x = x := by module example (h₁ : t - u = -(v - w)) (h₂ : t + u = v + w) : t = w := by - -- `linear_combination (norm := module) 2⁻¹ • h₁ + 2⁻¹ • h₂` - apply eq_of_eq (congr((2:K)⁻¹ • $h₁ + (2:K)⁻¹ • $h₂):) - module + linear_combination (norm := module) (2:K)⁻¹ • h₁ + (2:K)⁻¹ • h₂ end CharZeroField @@ -288,8 +268,7 @@ example (h₁ : 1 = a ^ 2 + b ^ 2) (h₂ : 1 - a ≠ 0) : ((2 / (1 - a)) ^ 2 * b ^ 2 + 4)⁻¹ • (4:K) • ((2 / (1 - a)) • y) + ((2 / (1 - a)) ^ 2 * b ^ 2 + 4)⁻¹ • ((2 / (1 - a)) ^ 2 * b ^ 2 - 4) • x = a • x + y := by - -- `linear_combination (norm := skip) (h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x)` - apply eq_of_eq (congr(($h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x)):) + linear_combination (norm := skip) (h₁ * (b ^ 2 + (1 - a) ^ 2)⁻¹) • (y + (a - 1) • x) match_scalars · field_simp ring From 38d4438a8695ba746705bb18a58eb1d104febfb8 Mon Sep 17 00:00:00 2001 From: Nick Ward <102917377+gio256@users.noreply.github.com> Date: Tue, 26 Nov 2024 19:59:18 +0000 Subject: [PATCH 193/250] chore(AlgebraicTopology): move Nerve.* to CategoryTheory.Nerve.* (#19521) Moves definitions to more appropriate namespaces: - Nerve.strictSegal -> CategoryTheory.Nerve.strictSegal - Nerve.quasicategory -> CategoryTheory.Nerve.quasicategory Co-Authored-By: [Emily Riehl](https://github.com/emilyriehl) --- Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean | 6 +++--- Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean | 6 ++---- 2 files changed, 5 insertions(+), 7 deletions(-) diff --git a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean index 22963dfe7b25f..5b8e8e517bd47 100644 --- a/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean +++ b/Mathlib/AlgebraicTopology/Quasicategory/Nerve.lean @@ -19,13 +19,13 @@ a quasicategory. universe v u -open CategoryTheory SSet +open SSet -namespace Nerve +namespace CategoryTheory.Nerve /-- By virtue of satisfying the `StrictSegal` condition, the nerve of a category is a `Quasicategory`. -/ instance quasicategory {C : Type u} [Category.{v} C] : Quasicategory (nerve C) := inferInstance -end Nerve +end CategoryTheory.Nerve diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean index 71e41b710ae25..7436832a79255 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/StrictSegal.lean @@ -156,12 +156,10 @@ end StrictSegal end SSet -namespace Nerve +namespace CategoryTheory.Nerve open SSet -variable {C : Type*} [Category C] {n : ℕ} - /-- Simplices in the nerve of categories are uniquely determined by their spine. Indeed, this property describes the essential image of the nerve functor.-/ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal (nerve C) where @@ -184,4 +182,4 @@ noncomputable instance strictSegal (C : Type u) [Category.{v} C] : StrictSegal ( · intro i hi apply ComposableArrows.mkOfObjOfMapSucc_map_succ -end Nerve +end CategoryTheory.Nerve From b25c29685772c4e73f377e6746ed44c2d3d29907 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 26 Nov 2024 20:08:00 +0000 Subject: [PATCH 194/250] chore(RingTheory/Localization): golfs and generalizations (#19310) + In RingTheory/LocalProperties/Basic: add some lemmas connecting localization of submodules and localization of ideals, and use results about the former to golf results about the latter. + In RingTheory/Localization/Module: golf a proof using the recently introduced IsLocalization.linearMap_compatibleSMul. + In Algebra/Module/Projective, Algebra/Module/LocalizedModule/IsLocalization, RingTheory/Localization/Submodule: generalize from AddCommGroup/CommRing to AddCommMonoid/CommSemiring. Also golf a lemma statement in the second file. + RingTheory/Localization/Algebra, RingTheory/Localization/Ideal: add two TODOs that are not done here for import reasons. --- .../LocalizedModule/IsLocalization.lean | 4 +- Mathlib/Algebra/Module/Projective.lean | 8 +- Mathlib/RingTheory/LocalProperties/Basic.lean | 82 ++++++++++--------- Mathlib/RingTheory/Localization/Algebra.lean | 1 + Mathlib/RingTheory/Localization/Ideal.lean | 1 + Mathlib/RingTheory/Localization/Module.lean | 17 +--- .../RingTheory/Localization/Submodule.lean | 20 +++-- 7 files changed, 65 insertions(+), 68 deletions(-) diff --git a/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean b/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean index fb79c4a5368da..aed527565f88b 100644 --- a/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean +++ b/Mathlib/Algebra/Module/LocalizedModule/IsLocalization.lean @@ -34,11 +34,11 @@ instance {A Aₛ} [CommSemiring A] [Algebra R A][CommSemiring Aₛ] [Algebra A A isLocalizedModule_iff_isLocalization.mpr h lemma isLocalizedModule_iff_isLocalization' (R') [CommSemiring R'] [Algebra R R'] : - IsLocalizedModule S (Algebra.ofId R R').toLinearMap ↔ IsLocalization S R' := by + IsLocalizedModule S (Algebra.linearMap R R') ↔ IsLocalization S R' := by convert isLocalizedModule_iff_isLocalization (S := S) (A := R) (Aₛ := R') exact (Submonoid.map_id S).symm -instance {A} [CommRing A] [Algebra R A] [IsLocalization S A] : +instance {A} [CommSemiring A] [Algebra R A] [IsLocalization S A] : IsLocalizedModule S (Algebra.linearMap R A) := (isLocalizedModule_iff_isLocalization' S _).mpr inferInstance diff --git a/Mathlib/Algebra/Module/Projective.lean b/Mathlib/Algebra/Module/Projective.lean index fdce19a9a6937..42ac6dfc4685e 100644 --- a/Mathlib/Algebra/Module/Projective.lean +++ b/Mathlib/Algebra/Module/Projective.lean @@ -194,13 +194,13 @@ end Semiring section Ring -variable {R : Type u} [Ring R] {P : Type v} [AddCommMonoid P] [Module R P] -variable {R₀ M N} [CommRing R₀] [Algebra R₀ R] [AddCommGroup M] [Module R₀ M] [Module R M] -variable [IsScalarTower R₀ R M] [AddCommGroup N] [Module R₀ N] +variable {R : Type u} [Semiring R] {P : Type v} [AddCommMonoid P] [Module R P] +variable {R₀ M N} [CommSemiring R₀] [Algebra R₀ R] [AddCommMonoid M] [Module R₀ M] [Module R M] +variable [IsScalarTower R₀ R M] [AddCommMonoid N] [Module R₀ N] /-- A module is projective iff it is the direct summand of a free module. -/ theorem Projective.iff_split : Module.Projective R P ↔ - ∃ (M : Type max u v) (_ : AddCommGroup M) (_ : Module R M) (_ : Module.Free R M) + ∃ (M : Type max u v) (_ : AddCommMonoid M) (_ : Module R M) (_ : Module.Free R M) (i : P →ₗ[R] M) (s : M →ₗ[R] P), s.comp i = LinearMap.id := ⟨fun ⟨i, hi⟩ ↦ ⟨P →₀ R, _, _, inferInstance, i, Finsupp.linearCombination R id, LinearMap.ext hi⟩, fun ⟨_, _, _, _, i, s, H⟩ ↦ Projective.of_split i s H⟩ diff --git a/Mathlib/RingTheory/LocalProperties/Basic.lean b/Mathlib/RingTheory/LocalProperties/Basic.lean index 0b332d9f25b8b..354847e24b104 100644 --- a/Mathlib/RingTheory/LocalProperties/Basic.lean +++ b/Mathlib/RingTheory/LocalProperties/Basic.lean @@ -3,10 +3,10 @@ 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.Ideal.Colon import Mathlib.RingTheory.Localization.AtPrime import Mathlib.RingTheory.Localization.BaseChange import Mathlib.RingTheory.Localization.Submodule +import Mathlib.RingTheory.LocalProperties.Submodule import Mathlib.RingTheory.RingHomProperties /-! @@ -45,12 +45,12 @@ open scoped Pointwise Classical universe u +section Properties + variable {R S : Type u} [CommRing R] [CommRing S] (f : R →+* S) variable (R' S' : Type u) [CommRing R'] [CommRing S'] variable [Algebra R R'] [Algebra S S'] -section Properties - section CommRing variable (P : ∀ (R : Type u) [CommRing R], Prop) @@ -396,7 +396,35 @@ end Properties section Ideal -open scoped nonZeroDivisors +variable {R : Type*} (S : Type*) [CommSemiring R] [CommSemiring S] [Algebra R S] +variable (p : Submonoid R) [IsLocalization p S] + +theorem Ideal.localized'_eq_map (I : Ideal R) : + haveI := (isLocalizedModule_iff_isLocalization' p S).mpr inferInstance + Submodule.localized' S p (Algebra.linearMap R S) I = I.map (algebraMap R S) := + SetLike.ext fun x ↦ by + simp_rw [Submodule.mem_localized', IsLocalization.mem_map_algebraMap_iff p, + IsLocalizedModule.mk'_eq_iff, mul_comm x, eq_comm (a := _ * x), ← Algebra.smul_def, + Prod.exists, Subtype.exists, ← exists_prop] + rfl + +theorem Ideal.localized₀_eq_restrictScalars_map (I : Ideal R) : + Submodule.localized₀ p (Algebra.linearMap R S) I = (I.map (algebraMap R S)).restrictScalars R := + congr(Submodule.restrictScalars R $(localized'_eq_map S p I)) + +theorem Algebra.idealMap_eq_ofEq_comp_toLocalized₀ (I : Ideal R) : + haveI := (isLocalizedModule_iff_isLocalization' p S).mpr inferInstance + Algebra.idealMap S I = + (LinearEquiv.ofEq _ _ <| Ideal.localized₀_eq_restrictScalars_map S p I).toLinearMap ∘ₗ + Submodule.toLocalized₀ p (Algebra.linearMap R S) I := + rfl + +theorem Ideal.mem_of_localization_maximal {r : R} {J : Ideal R} + (h : ∀ (P : Ideal R) (_ : P.IsMaximal), + algebraMap R _ r ∈ Ideal.map (algebraMap R (Localization.AtPrime P)) J) : + r ∈ J := + Submodule.mem_of_localization_maximal _ _ _ _ fun P hP ↦ by + apply (localized'_eq_map (Localization.AtPrime P) P.primeCompl J).symm ▸ h P hP /-- Let `I J : Ideal R`. If the localization of `I` at each maximal ideal `P` is included in the localization of `J` at `P`, then `I ≤ J`. -/ @@ -404,22 +432,8 @@ theorem Ideal.le_of_localization_maximal {I J : Ideal R} (h : ∀ (P : Ideal R) (_ : P.IsMaximal), Ideal.map (algebraMap R (Localization.AtPrime P)) I ≤ Ideal.map (algebraMap R (Localization.AtPrime P)) J) : - I ≤ J := by - intro x hx - suffices J.colon (Ideal.span {x}) = ⊤ by - simpa using Submodule.mem_colon.mp - (show (1 : R) ∈ J.colon (Ideal.span {x}) from this.symm ▸ Submodule.mem_top) x - (Ideal.mem_span_singleton_self x) - refine Not.imp_symm (J.colon (Ideal.span {x})).exists_le_maximal ?_ - push_neg - intro P hP le - obtain ⟨⟨⟨a, ha⟩, ⟨s, hs⟩⟩, eq⟩ := - (IsLocalization.mem_map_algebraMap_iff P.primeCompl _).mp (h P hP (Ideal.mem_map_of_mem _ hx)) - rw [← _root_.map_mul, ← sub_eq_zero, ← map_sub] at eq - obtain ⟨⟨m, hm⟩, eq⟩ := (IsLocalization.map_eq_zero_iff P.primeCompl _ _).mp eq - refine hs ((hP.isPrime.mem_or_mem (le (Ideal.mem_colon_singleton.mpr ?_))).resolve_right hm) - simp only [Subtype.coe_mk, mul_sub, sub_eq_zero, mul_comm x s, mul_left_comm] at eq - simpa only [mul_assoc, eq] using J.mul_mem_left m ha + I ≤ J := + fun _ hm ↦ mem_of_localization_maximal fun P hP ↦ h P hP (mem_map_of_mem _ hm) /-- Let `I J : Ideal R`. If the localization of `I` at each maximal ideal `P` is equal to the localization of `J` at `P`, then `I = J`. -/ @@ -428,8 +442,8 @@ theorem Ideal.eq_of_localization_maximal {I J : Ideal R} Ideal.map (algebraMap R (Localization.AtPrime P)) I = Ideal.map (algebraMap R (Localization.AtPrime P)) J) : I = J := - le_antisymm (Ideal.le_of_localization_maximal fun P hP => (h P hP).le) - (Ideal.le_of_localization_maximal fun P hP => (h P hP).ge) + le_antisymm (le_of_localization_maximal fun P hP ↦ (h P hP).le) + (le_of_localization_maximal fun P hP ↦ (h P hP).ge) /-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ theorem ideal_eq_bot_of_localization' (I : Ideal R) @@ -438,28 +452,16 @@ theorem ideal_eq_bot_of_localization' (I : Ideal R) I = ⊥ := Ideal.eq_of_localization_maximal fun P hP => by simpa using h P hP --- TODO: This proof should work for all modules, once we have enough material on submodules of --- localized modules. +theorem eq_zero_of_localization (r : R) + (h : ∀ (J : Ideal R) (_ : J.IsMaximal), algebraMap R (Localization.AtPrime J) r = 0) : + r = 0 := + Module.eq_zero_of_localization_maximal _ (fun _ _ ↦ Algebra.linearMap R _) r h + /-- An ideal is trivial if its localization at every maximal ideal is trivial. -/ theorem ideal_eq_bot_of_localization (I : Ideal R) (h : ∀ (J : Ideal R) (_ : J.IsMaximal), IsLocalization.coeSubmodule (Localization.AtPrime J) I = ⊥) : I = ⊥ := - ideal_eq_bot_of_localization' _ fun P hP => - (Ideal.map_eq_bot_iff_le_ker _).mpr fun x hx => by - rw [RingHom.mem_ker, ← Submodule.mem_bot R, ← h P hP, IsLocalization.mem_coeSubmodule] - exact ⟨x, hx, rfl⟩ - -theorem eq_zero_of_localization (r : R) - (h : ∀ (J : Ideal R) (_ : J.IsMaximal), algebraMap R (Localization.AtPrime J) r = 0) : - r = 0 := by - rw [← Ideal.span_singleton_eq_bot] - apply ideal_eq_bot_of_localization - intro J hJ - delta IsLocalization.coeSubmodule - erw [Submodule.map_span, Submodule.span_eq_bot] - rintro _ ⟨_, h', rfl⟩ - cases Set.mem_singleton_iff.mpr h' - exact h J hJ + bot_unique fun r hr ↦ eq_zero_of_localization r fun J hJ ↦ (h J hJ).le ⟨r, hr, rfl⟩ end Ideal diff --git a/Mathlib/RingTheory/Localization/Algebra.lean b/Mathlib/RingTheory/Localization/Algebra.lean index 0210fb6bc2be8..b52b401bdd5bf 100644 --- a/Mathlib/RingTheory/Localization/Algebra.lean +++ b/Mathlib/RingTheory/Localization/Algebra.lean @@ -32,6 +32,7 @@ variable {R S P : Type*} (Q : Type*) [CommSemiring R] [CommSemiring S] [CommSemi open IsLocalization in variable (M S) in /-- The span of `I` in a localization of `R` at `M` is the localization of `I` at `M`. -/ +-- TODO: golf using `Ideal.localized'_eq_map` instance Algebra.idealMap_isLocalizedModule (I : Ideal R) : IsLocalizedModule M (Algebra.idealMap I (S := S)) where map_units x := diff --git a/Mathlib/RingTheory/Localization/Ideal.lean b/Mathlib/RingTheory/Localization/Ideal.lean index 4d5fd06147aed..126d2cbf0c729 100644 --- a/Mathlib/RingTheory/Localization/Ideal.lean +++ b/Mathlib/RingTheory/Localization/Ideal.lean @@ -28,6 +28,7 @@ variable [Algebra R S] [IsLocalization M S] In practice, this ideal differs only in that the carrier set is defined explicitly. This definition is only meant to be used in proving `mem_map_algebraMap_iff`, and any proof that needs to refer to the explicit carrier set should use that theorem. -/ +-- TODO: golf this using `Submodule.localized'` private def map_ideal (I : Ideal R) : Ideal S where carrier := { z : S | ∃ x : I × M, z * algebraMap R S x.2 = algebraMap R S x.1 } zero_mem' := ⟨⟨0, 1⟩, by simp⟩ diff --git a/Mathlib/RingTheory/Localization/Module.lean b/Mathlib/RingTheory/Localization/Module.lean index 328a15ccb228b..40c9dccaad408 100644 --- a/Mathlib/RingTheory/Localization/Module.lean +++ b/Mathlib/RingTheory/Localization/Module.lean @@ -207,17 +207,7 @@ open IsLocalization def LinearMap.extendScalarsOfIsLocalization (f : M →ₗ[R] N) : M →ₗ[A] N where toFun := f map_add' := f.map_add - map_smul' := by - intro r m - simp only [RingHom.id_apply] - rcases mk'_surjective S r with ⟨r, s, rfl⟩ - calc f (mk' A r s • m) - = ((s : R) • mk' A 1 s) • f (mk' A r s • m) := by simp - _ = (mk' A 1 s) • (s : R) • f (mk' A r s • m) := by rw [smul_comm, smul_assoc] - _ = (mk' A 1 s) • f ((s : R) • mk' A r s • m) := by simp - _ = (mk' A 1 s) • f (r • m) := by rw [← smul_assoc, smul_mk'_self, algebraMap_smul] - _ = (mk' A 1 s) • r • f m := by simp - _ = mk' A r s • f m := by rw [smul_comm, ← smul_assoc, smul_mk'_one] + map_smul' := (IsLocalization.linearMap_compatibleSMul S A M N).map_smul _ @[simp] lemma LinearMap.restrictScalars_extendScalarsOfIsLocalization (f : M →ₗ[R] N) : (f.extendScalarsOfIsLocalization S A).restrictScalars R = f := rfl @@ -257,8 +247,7 @@ variable [IsScalarTower R Rₛ M'] [IsScalarTower R Rₛ N'] [IsLocalization S R @[simps!] noncomputable def mapExtendScalars : (M →ₗ[R] N) →ₗ[R] (M' →ₗ[Rₛ] N') := - ((LinearMap.extendScalarsOfIsLocalizationEquiv - S Rₛ).restrictScalars R).toLinearMap.comp (map S f g) + ((LinearMap.extendScalarsOfIsLocalizationEquiv S Rₛ).restrictScalars R).toLinearMap ∘ₗ map S f g end IsLocalizedModule @@ -273,7 +262,7 @@ noncomputable def LocalizedModule.map : (M →ₗ[R] N) →ₗ[R] (LocalizedModule S M →ₗ[Localization S] LocalizedModule S N) := IsLocalizedModule.mapExtendScalars S (LocalizedModule.mkLinearMap S M) - (LocalizedModule.mkLinearMap S N) (Localization S) + (LocalizedModule.mkLinearMap S N) (Localization S) @[simp] lemma LocalizedModule.map_mk (f : M →ₗ[R] N) (x y) : diff --git a/Mathlib/RingTheory/Localization/Submodule.lean b/Mathlib/RingTheory/Localization/Submodule.lean index ec8df59b713d2..a1e3a22a8e03c 100644 --- a/Mathlib/RingTheory/Localization/Submodule.lean +++ b/Mathlib/RingTheory/Localization/Submodule.lean @@ -20,7 +20,7 @@ commutative ring, field of fractions -/ -variable {R : Type*} [CommRing R] (M : Submonoid R) (S : Type*) [CommRing S] +variable {R : Type*} [CommSemiring R] (M : Submonoid R) (S : Type*) [CommSemiring S] variable [Algebra R S] namespace IsLocalization @@ -59,7 +59,7 @@ theorem coeSubmodule_mul (I J : Ideal R) : theorem coeSubmodule_fg (hS : Function.Injective (algebraMap R S)) (I : Ideal R) : Submodule.FG (coeSubmodule S I) ↔ Submodule.FG I := - ⟨Submodule.fg_of_fg_map _ (LinearMap.ker_eq_bot.mpr hS), Submodule.FG.map _⟩ + ⟨Submodule.fg_of_fg_map_injective _ hS, Submodule.FG.map _⟩ @[simp] theorem coeSubmodule_span (s : Set R) : @@ -78,7 +78,10 @@ theorem isNoetherianRing (h : IsNoetherianRing R) : IsNoetherianRing S := by rw [isNoetherianRing_iff, isNoetherian_iff] at h ⊢ exact OrderEmbedding.wellFounded (IsLocalization.orderEmbedding M S).dual h -variable {S M} +section NonZeroDivisors + +variable {R : Type*} [CommRing R] {M : Submonoid R} + {S : Type*} [CommRing S] [Algebra R S] [IsLocalization M S] @[mono] theorem coeSubmodule_le_coeSubmodule (h : M ≤ nonZeroDivisors R) {I J : Ideal R} : @@ -87,7 +90,6 @@ theorem coeSubmodule_le_coeSubmodule (h : M ≤ nonZeroDivisors R) {I J : Ideal Submodule.map_le_map_iff_of_injective (f := Algebra.linearMap R S) (IsLocalization.injective _ h) _ _ - @[mono] theorem coeSubmodule_strictMono (h : M ≤ nonZeroDivisors R) : StrictMono (coeSubmodule S : Ideal R → Submodule R S) := @@ -109,9 +111,11 @@ theorem coeSubmodule_isPrincipal {I : Ideal R} (h : M ≤ nonZeroDivisors R) : · refine ⟨⟨algebraMap R S x, ?_⟩⟩ rw [hx, Ideal.submodule_span_eq, coeSubmodule_span_singleton] -variable {S} (M) +end NonZeroDivisors + +variable {S} -theorem mem_span_iff {N : Type*} [AddCommGroup N] [Module R N] [Module S N] [IsScalarTower R S N] +theorem mem_span_iff {N : Type*} [AddCommMonoid N] [Module R N] [Module S N] [IsScalarTower R S N] {x : N} {a : Set N} : x ∈ Submodule.span S a ↔ ∃ y ∈ Submodule.span R a, ∃ z : M, x = mk' S 1 z • y := by constructor @@ -155,11 +159,11 @@ namespace IsFractionRing open IsLocalization -variable {K : Type*} +variable {R K : Type*} section CommRing -variable [CommRing K] [Algebra R K] [IsFractionRing R K] +variable [CommRing R] [CommRing K] [Algebra R K] [IsFractionRing R K] @[simp, mono] theorem coeSubmodule_le_coeSubmodule {I J : Ideal R} : From c8c981908a997925fb9110b09da9f0d13a02b808 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Tue, 26 Nov 2024 22:10:33 +0000 Subject: [PATCH 195/250] fix: build cache issues (#19524) Fixes compatibility issues between Cache and Lake's builtin build cache. --- .github/build.in.yml | 9 +++++++++ .github/workflows/bors.yml | 9 +++++++++ .github/workflows/build.yml | 9 +++++++++ .github/workflows/build_fork.yml | 9 +++++++++ Cache/IO.lean | 3 ++- Cache/Requests.lean | 14 +++++++++++++- 6 files changed, 51 insertions(+), 2 deletions(-) diff --git a/.github/build.in.yml b/.github/build.in.yml index bab96d693f6b2..141b5316926ab 100644 --- a/.github/build.in.yml +++ b/.github/build.in.yml @@ -4,6 +4,14 @@ ### NB: and regenerate those files by manually running ### NB: .github/workflows/mk_build_yml.sh +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -121,6 +129,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/bors.yml b/.github/workflows/bors.yml index e59f48d2c7267..3bd02efbb94dd 100644 --- a/.github/workflows/bors.yml +++ b/.github/workflows/bors.yml @@ -14,6 +14,14 @@ on: name: continuous integration (staging) +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -131,6 +139,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/build.yml b/.github/workflows/build.yml index 1c778a5720c23..b5f11a1ea1f03 100644 --- a/.github/workflows/build.yml +++ b/.github/workflows/build.yml @@ -21,6 +21,14 @@ on: name: continuous integration +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -138,6 +146,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/.github/workflows/build_fork.yml b/.github/workflows/build_fork.yml index e92f87fd50a68..4da591010b8b7 100644 --- a/.github/workflows/build_fork.yml +++ b/.github/workflows/build_fork.yml @@ -18,6 +18,14 @@ on: name: continuous integration (mathlib forks) +env: + # Disable Lake's automatic fetching of cloud builds. + # Lake's cache is currently incompatible with Mathlib's `lake exe cache`. + # This is because Mathlib's Cache assumes all build aritfacts present in the build directory + # are valid by-products of the Mathlib build. Build artifacts fetched from Lake's cache do + # not necessarily satisfy this property. + LAKE_NO_CACHE: true + jobs: # Cancels previous runs of jobs in this file cancel: @@ -135,6 +143,7 @@ jobs: - name: prune ProofWidgets .lake run: | + lake build proofwidgets:release # The ProofWidgets release contains not just the `.js` (which we need in order to build) # but also `.oleans`, which may have been built with the wrong toolchain. # This removes them. diff --git a/Cache/IO.lean b/Cache/IO.lean index 522364f147d94..358683a890f1c 100644 --- a/Cache/IO.lean +++ b/Cache/IO.lean @@ -92,6 +92,7 @@ abbrev PackageDirs := Lean.RBMap String FilePath compare structure CacheM.Context where mathlibDepPath : FilePath packageDirs : PackageDirs + proofWidgetsBuildDir : FilePath abbrev CacheM := ReaderT CacheM.Context IO @@ -141,7 +142,7 @@ private def CacheM.getContext : IO CacheM.Context := do ("ImportGraph", LAKEPACKAGESDIR / "importGraph"), ("LeanSearchClient", LAKEPACKAGESDIR / "LeanSearchClient"), ("Plausible", LAKEPACKAGESDIR / "plausible") - ]⟩ + ], LAKEPACKAGESDIR / "proofwidgets" / ".lake" / "build"⟩ def CacheM.run (f : CacheM α) : IO α := do ReaderT.run f (← getContext) diff --git a/Cache/Requests.lean b/Cache/Requests.lean index 71b74cb4b051d..295d65cd96693 100644 --- a/Cache/Requests.lean +++ b/Cache/Requests.lean @@ -159,11 +159,23 @@ into the `lean-toolchain` file at the root directory of your project" IO.Process.exit 1 return () +/-- Fetches the ProofWidgets cloud release and prunes non-JS files. -/ +def getProofWidgets (buildDir : FilePath) : IO Unit := do + if (← buildDir.pathExists) then return + -- Unpack the ProofWidgets cloud release (for its `.js` files) + let exitCode ← (← IO.Process.spawn {cmd := "lake", args := #["-q", "build", "proofwidgets:release"]}).wait + if exitCode != 0 then + throw <| IO.userError s!"Failed to fetch ProofWidgets cloud release: lake failed with error code {exitCode}" + -- prune non-js ProofWidgets files (e.g., `olean`, `.c`) + IO.FS.removeDirAll (buildDir / "lib") + IO.FS.removeDirAll (buildDir / "ir") + /-- Downloads missing files, and unpacks files. -/ def getFiles (hashMap : IO.HashMap) (forceDownload forceUnpack parallel decompress : Bool) : IO.CacheM Unit := do let isMathlibRoot ← IO.isMathlibRoot - if !isMathlibRoot then checkForToolchainMismatch + unless isMathlibRoot do checkForToolchainMismatch + getProofWidgets (← read).proofWidgetsBuildDir downloadFiles hashMap forceDownload parallel if decompress then IO.unpackCache hashMap forceUnpack From 4ecf030d4de9171ac19aeefd9273f367675687b0 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 09:12:15 +1100 Subject: [PATCH 196/250] fix --- Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean index 876f75a97121b..30441b7dda986 100644 --- a/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Finset/Basic.lean @@ -958,7 +958,7 @@ protected lemma pow_right_monotone (hs : 1 ∈ s) : Monotone (s ^ ·) := pow_right_monotone <| one_subset.2 hs @[to_additive (attr := gcongr)] -lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := pow_left_mono _ hst +lemma pow_subset_pow_left (hst : s ⊆ t) : s ^ n ⊆ t ^ n := subset_of_le (pow_left_mono n hst) @[to_additive (attr := gcongr)] lemma pow_subset_pow_right (hs : 1 ∈ s) (hmn : m ≤ n) : s ^ m ⊆ s ^ n := @@ -975,7 +975,7 @@ alias nsmul_subset_nsmul_of_zero_mem := nsmul_subset_nsmul_right @[to_additive] lemma pow_subset_pow_mul_of_sq_subset_mul (hst : s ^ 2 ⊆ t * s) (hn : n ≠ 0) : - s ^ n ⊆ t ^ (n - 1) * s := pow_le_pow_mul_of_sq_le_mul hst hn + s ^ n ⊆ t ^ (n - 1) * s := subset_of_le (pow_le_pow_mul_of_sq_le_mul hst hn) @[to_additive (attr := simp) nsmul_empty] lemma empty_pow (hn : n ≠ 0) : (∅ : Finset α) ^ n = ∅ := match n with | n + 1 => by simp [pow_succ] From 21af8545cb9daf04f22cddc30a99fd7dbbf3543b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 26 Nov 2024 22:19:05 +0000 Subject: [PATCH 197/250] chore: adaptation to batteries#1055 (#19487) Co-authored-by: F. G. Dorais Co-authored-by: Kim Morrison Co-authored-by: Matthew Ballard --- Mathlib/Data/Fin/VecNotation.lean | 1 - Mathlib/Data/List/Defs.lean | 8 -------- Mathlib/Data/List/FinRange.lean | 18 +----------------- Mathlib/ModelTheory/Encoding.lean | 6 +++--- lake-manifest.json | 2 +- lakefile.lean | 3 +-- 6 files changed, 6 insertions(+), 32 deletions(-) diff --git a/Mathlib/Data/Fin/VecNotation.lean b/Mathlib/Data/Fin/VecNotation.lean index 9d31131cdaf6e..1b8ca2379a123 100644 --- a/Mathlib/Data/Fin/VecNotation.lean +++ b/Mathlib/Data/Fin/VecNotation.lean @@ -4,7 +4,6 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen -/ import Mathlib.Data.Fin.Tuple.Basic -import Mathlib.Data.List.Defs /-! # Matrix and vector notation diff --git a/Mathlib/Data/List/Defs.lean b/Mathlib/Data/List/Defs.lean index 4b86b7218cdfb..b340e9b270a42 100644 --- a/Mathlib/Data/List/Defs.lean +++ b/Mathlib/Data/List/Defs.lean @@ -492,14 +492,6 @@ end MapAccumr set_option allowUnsafeReducibility true in attribute [semireducible] Fin.foldr.loop -/-- All elements of `Fin n`, from `0` to `n-1`. The corresponding finset is `Finset.univ`. -/ --- Note that we use `ofFn (fun x => x)` instead of `ofFn id` to avoid leaving `id` in the terms --- for e.g. `Fintype (Fin n)`. -def finRange (n : Nat) : List (Fin n) := ofFn (fun x => x) - --- Verify that `finRange` is semireducible. -example : finRange 3 = [0, 1, 2] := rfl - section Deprecated @[deprecated List.mem_cons (since := "2024-08-10")] diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index bf592cfe78937..83ccab9a12979 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -5,6 +5,7 @@ Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn import Batteries.Data.List.Perm +import Batteries.Data.List.FinRange import Mathlib.Data.List.Nodup /-! @@ -25,9 +26,6 @@ variable {α : Type u} theorem finRange_eq_pmap_range (n : ℕ) : finRange n = (range n).pmap Fin.mk (by simp) := by apply List.ext_getElem <;> simp [finRange] -@[simp] -theorem finRange_zero : finRange 0 = [] := rfl - @[simp] theorem mem_finRange {n : ℕ} (a : Fin n) : a ∈ finRange n := by rw [finRange_eq_pmap_range] @@ -40,10 +38,6 @@ theorem nodup_finRange (n : ℕ) : (finRange n).Nodup := by rw [finRange_eq_pmap_range] exact (Pairwise.pmap (nodup_range n) _) fun _ _ _ _ => @Fin.ne_of_val_ne _ ⟨_, _⟩ ⟨_, _⟩ -@[simp] -theorem length_finRange (n : ℕ) : (finRange n).length = n := by - simp [finRange] - @[simp] theorem finRange_eq_nil {n : ℕ} : finRange n = [] ↔ n = 0 := by rw [← length_eq_zero, length_finRange] @@ -56,11 +50,6 @@ theorem pairwise_le_finRange (n : ℕ) : Pairwise (· ≤ ·) (finRange n) := by rw [finRange_eq_pmap_range] exact (List.pairwise_le_range n).pmap (by simp) (by simp) -@[simp] -theorem getElem_finRange {n : ℕ} {i : ℕ} (h) : - (finRange n)[i] = ⟨i, length_finRange n ▸ h⟩ := by - simp [finRange, getElem_range, getElem_pmap] - -- Porting note (https://github.com/leanprover-community/mathlib4/issues/10756): new theorem theorem get_finRange {n : ℕ} {i : ℕ} (h) : (finRange n).get ⟨i, h⟩ = ⟨i, length_finRange n ▸ h⟩ := by @@ -88,11 +77,6 @@ theorem finRange_succ_eq_map (n : ℕ) : finRange n.succ = 0 :: (finRange n).map map_map] simp only [Function.comp_def, Fin.val_succ] -theorem finRange_succ (n : ℕ) : - finRange n.succ = (finRange n |>.map Fin.castSucc |>.concat (.last _)) := by - apply map_injective_iff.mpr Fin.val_injective - simp [range_succ, Function.comp_def] - -- Porting note: `map_nth_le` moved to `List.finRange_map_get` in Data.List.Range theorem ofFn_eq_pmap {n} {f : Fin n → α} : diff --git a/Mathlib/ModelTheory/Encoding.lean b/Mathlib/ModelTheory/Encoding.lean index 180cb86b870ed..e3832e469136b 100644 --- a/Mathlib/ModelTheory/Encoding.lean +++ b/Mathlib/ModelTheory/Encoding.lean @@ -82,8 +82,8 @@ theorem listDecode_encode_list (l : List (L.Term α)) : simp only [h, length_append, length_map, length_finRange, le_add_iff_nonneg_right, _root_.zero_le, ↓reduceDIte, getElem_fin, cons.injEq, func.injEq, heq_eq_eq, true_and] refine ⟨funext (fun i => ?_), ?_⟩ - · rw [List.getElem_append_left, List.getElem_map, List.getElem_finRange] - simp only [length_map, length_finRange, i.2] + · simp only [length_map, length_finRange, is_lt, getElem_append_left, getElem_map, + getElem_finRange, cast_mk, Fin.eta] · simp only [length_map, length_finRange, drop_left'] /-- An encoding of terms as lists. -/ @@ -237,7 +237,7 @@ theorem listDecode_encode_list (l : List (Σn, L.BoundedFormula α n)) : get?_eq_some, length_append, length_map, length_finRange] refine ⟨lt_of_lt_of_le i.2 le_self_add, ?_⟩ rw [get_eq_getElem, getElem_append_left, getElem_map] - · simp only [getElem_finRange, Fin.eta, Function.comp_apply, Sum.getLeft?] + · simp only [getElem_finRange, cast_mk, Fin.eta, Function.comp_apply, Sum.getLeft?_inl] · simp only [length_map, length_finRange, is_lt] rw [dif_pos] swap diff --git a/lake-manifest.json b/lake-manifest.json index 8152b35704a8a..a3160f181a6cb 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "7488499a8aad6ffada87ab6db73673d88dc04c97", + "rev": "f6d16c2a073e0385a362ad3e5d9e7e8260e298d6", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", diff --git a/lakefile.lean b/lakefile.lean index 300ff2124f428..cab011aba5693 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -23,8 +23,7 @@ require "leanprover-community" / "plausible" @ git "main" /-- These options are used * as `leanOptions`, prefixed by `` `weak``, so that `lake build` uses them; * as `moreServerArgs`, to set their default value in mathlib - (as well as `Archive`, `Counterexamples` and `test`). --/ + (as well as `Archive`, `Counterexamples` and `test`). -/ abbrev mathlibOnlyLinters : Array LeanOption := #[ ⟨`linter.docPrime, true⟩, ⟨`linter.hashCommand, true⟩, From 337de2f56a70c9dcfa3f98bfa006823f42c18919 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 09:21:05 +1100 Subject: [PATCH 198/250] fix --- Mathlib/Algebra/Category/Ring/Constructions.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index ca21d51022790..6a11059149823 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -106,8 +106,8 @@ def pushoutCoconeIsColimit : Limits.IsColimit (pushoutCocone R A B) := lemma isPushout_tensorProduct (R A B : Type u) [CommRing R] [CommRing A] [CommRing B] [Algebra R A] [Algebra R B] : IsPushout (ofHom <| algebraMap R A) (ofHom <| algebraMap R B) - (ofHom <| Algebra.TensorProduct.includeLeftRingHom (R := R)) - (ofHom <| Algebra.TensorProduct.includeRight.toRingHom (R := R)) where + (ofHom (S := A ⊗[R] B) <| Algebra.TensorProduct.includeLeftRingHom) + (ofHom (S := A ⊗[R] B) <| Algebra.TensorProduct.includeRight.toRingHom) where w := by ext simp From 30c0e01daad4ff6def8c2d5dbc28b0cccd5d250e Mon Sep 17 00:00:00 2001 From: Daniel Carranza Date: Tue, 26 Nov 2024 22:40:37 +0000 Subject: [PATCH 199/250] feat(CategoryTheory/Enriched/Opposite): the opposite category of an enriched category (#19483) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit For a `V`-category `C`, define the opposite `V`-category as an instance on the type `Cᵒᵖ`, construct an equivalence between the underlying category `ForgetEnrichment V (Cᵒᵖ)` and the opposite category `(ForgetEnrichment V C)ᵒᵖ`, and construct an instance of `EnrichedOrdinaryCategory V (Cᵒᵖ)` in the case that `C` comes with an instance of `EnrichedOrdinaryCategory V C`. Co-authored-by: Daniel Carranza <61289743+daniel-carranza@users.noreply.github.com> --- Mathlib.lean | 1 + Mathlib/CategoryTheory/Enriched/Opposite.lean | 132 ++++++++++++++++++ 2 files changed, 133 insertions(+) create mode 100644 Mathlib/CategoryTheory/Enriched/Opposite.lean diff --git a/Mathlib.lean b/Mathlib.lean index 3741c55a437ae..c971dda66ba49 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1682,6 +1682,7 @@ import Mathlib.CategoryTheory.Endofunctor.Algebra import Mathlib.CategoryTheory.Endomorphism import Mathlib.CategoryTheory.Enriched.Basic import Mathlib.CategoryTheory.Enriched.FunctorCategory +import Mathlib.CategoryTheory.Enriched.Opposite import Mathlib.CategoryTheory.Enriched.Ordinary import Mathlib.CategoryTheory.EpiMono import Mathlib.CategoryTheory.EqToHom diff --git a/Mathlib/CategoryTheory/Enriched/Opposite.lean b/Mathlib/CategoryTheory/Enriched/Opposite.lean new file mode 100644 index 0000000000000..d25af18fc5ba9 --- /dev/null +++ b/Mathlib/CategoryTheory/Enriched/Opposite.lean @@ -0,0 +1,132 @@ +/- +Copyright (c) 2024 Daniel Carranza. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Daniel Carranza +-/ +import Mathlib.CategoryTheory.Enriched.Ordinary +import Mathlib.CategoryTheory.Monoidal.Braided.Basic + +/-! + +# The opposite category of an enriched category + +When a monoidal category `V` is braided, we may define the opposite `V`-category of a +`V`-category. The symmetry map is required to define the composition morphism. + +This file constructs the opposite `V`-category as an instance on the type `Cᵒᵖ` and constructs an +equivalence between + • `ForgetEnrichment V (Cᵒᵖ)`, the underlying category of the `V`-category `Cᵒᵖ`; and + • `(ForgetEnrichment V C)ᵒᵖ`, the opposite category of the underlying category of `C`. +We also show that if `C` is an enriched ordinary category (i.e. a category enriched in `V` +equipped with an identification `(X ⟶ Y) ≃ (𝟙_ V ⟶ (X ⟶[V] Y))`) then `Cᵒᵖ` is again +an enriched ordinary category. + +-/ + +universe v₁ u₁ v u + +namespace CategoryTheory + +open MonoidalCategory BraidedCategory + +variable (V : Type u₁) [Category.{v₁} V] [MonoidalCategory V] [BraidedCategory V] + +section + +variable (C : Type u) [EnrichedCategory V C] + +/-- For a `V`-category `C`, construct the opposite `V`-category structure on the type `Cᵒᵖ` +using the braiding in `V`. -/ +instance EnrichedCategory.opposite : EnrichedCategory V Cᵒᵖ where + Hom y x := EnrichedCategory.Hom x.unop y.unop + id x := EnrichedCategory.id x.unop + comp z y x := (β_ _ _).hom ≫ EnrichedCategory.comp (x.unop) (y.unop) (z.unop) + id_comp _ _ := by + simp only [braiding_naturality_left_assoc, braiding_tensorUnit_left, + Category.assoc, Iso.inv_hom_id_assoc] + exact EnrichedCategory.comp_id _ _ + comp_id _ _ := by + simp only [braiding_naturality_right_assoc, braiding_tensorUnit_right, + Category.assoc, Iso.inv_hom_id_assoc] + exact EnrichedCategory.id_comp _ _ + assoc _ _ _ _ := by + simp only [braiding_naturality_left_assoc, + MonoidalCategory.whiskerLeft_comp, Category.assoc] + rw [← EnrichedCategory.assoc] + simp only [braiding_tensor_left, Category.assoc, Iso.inv_hom_id_assoc, + braiding_naturality_right_assoc, braiding_tensor_right] + +end + +/-- Unfold the definition of composition in the enriched opposite category. -/ +@[reassoc] +lemma eComp_op_eq {C : Type u} [EnrichedCategory V C] (x y z : Cᵒᵖ) : + eComp V z y x = (β_ _ _).hom ≫ eComp V x.unop y.unop z.unop := + rfl + +/-- When composing a tensor product of morphisms with the `V`-composition morphism in `Cᵒᵖ`, +this re-writes the `V`-composition to be in `C` and moves the braiding to the left. -/ +@[reassoc] +lemma tensorHom_eComp_op_eq {C : Type u} [EnrichedCategory V C] {x y z : Cᵒᵖ} {v w : V} + (f : v ⟶ EnrichedCategory.Hom z y) (g : w ⟶ EnrichedCategory.Hom y x) : + (f ⊗ g) ≫ eComp V z y x = (β_ v w).hom ≫ (g ⊗ f) ≫ eComp V x.unop y.unop z.unop := by + rw [eComp_op_eq] + exact braiding_naturality_assoc f g _ + +-- This section establishes the equivalence on underlying categories +section + +open ForgetEnrichment + +variable (C : Type u) [EnrichedCategory V C] + +/-- The functor going from the underlying category of the enriched category `Cᵒᵖ` +to the opposite of the underlying category of the enriched category `C`. -/ +def forgetEnrichmentOppositeEquivalence.functor : + ForgetEnrichment V Cᵒᵖ ⥤ (ForgetEnrichment V C)ᵒᵖ where + obj x := x + map {x y} f := f.op + map_comp {x y z} f g := by + have : (f ≫ g) = homTo V (f ≫ g) := rfl + rw [this, forgetEnrichment_comp, Category.assoc, tensorHom_eComp_op_eq, + leftUnitor_inv_braiding_assoc, ← unitors_inv_equal, ← Category.assoc] + congr 1 + +/-- The functor going from the opposite of the underlying category of the enriched category `C` +to the underlying category of the enriched category `Cᵒᵖ`. -/ +def forgetEnrichmentOppositeEquivalence.inverse : + (ForgetEnrichment V C)ᵒᵖ ⥤ ForgetEnrichment V Cᵒᵖ where + obj x := x + map {x y} f := f.unop + map_comp {x y z} f g := by + have : g.unop ≫ f.unop = homTo V (g.unop ≫ f.unop) := rfl + dsimp + rw [this, forgetEnrichment_comp, Category.assoc, unitors_inv_equal, + ← leftUnitor_inv_braiding_assoc] + have : (β_ _ _).hom ≫ (homTo V g.unop ⊗ homTo V f.unop) ≫ + eComp V («to» V z.unop) («to» V y.unop) («to» V x.unop) = + ((homTo V f.unop) ⊗ (homTo V g.unop)) ≫ eComp V x y z := (tensorHom_eComp_op_eq V _ _).symm + rw [this, ← Category.assoc] + congr 1 + +/-- The equivalence between the underlying category of the enriched category `Cᵒᵖ` and +the opposite of the underlying category of the enriched category `C`. -/ +@[simps] +def forgetEnrichmentOppositeEquivalence : ForgetEnrichment V Cᵒᵖ ≌ (ForgetEnrichment V C)ᵒᵖ where + functor := forgetEnrichmentOppositeEquivalence.functor V C + inverse := forgetEnrichmentOppositeEquivalence.inverse V C + unitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) + counitIso := NatIso.ofComponents (fun _ ↦ Iso.refl _) + +/-- If `D` is an enriched ordinary category then `Dᵒᵖ` is an enriched ordinary category. -/ +instance EnrichedOrdinaryCategory.opposite {D : Type u} [Category.{v} D] + [EnrichedOrdinaryCategory V D] : EnrichedOrdinaryCategory V Dᵒᵖ where + homEquiv := Quiver.Hom.opEquiv.symm.trans homEquiv + homEquiv_id x := homEquiv_id (x.unop) + homEquiv_comp f g := by + simp only [unop_comp, tensorHom_eComp_op_eq, leftUnitor_inv_braiding_assoc, ← unitors_inv_equal] + exact homEquiv_comp g.unop f.unop + +end + +end CategoryTheory From 095a1c66b3101c42d97d114598a98d7e90916865 Mon Sep 17 00:00:00 2001 From: Junyan Xu Date: Tue, 26 Nov 2024 22:40:38 +0000 Subject: [PATCH 200/250] feat: semiring with discrete PrimeSpectrum (#19496) If the prime spectrum of a commutative semiring R has discrete Zariski topology, then R is canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. This was previously only proven for commutative rings R; a different proof using the sheaf property is required to generalize it. --- .../GammaSpecAdjunction.lean | 2 +- .../PrimeSpectrum/Basic.lean | 84 +++++++------ Mathlib/RingTheory/Ideal/Basic.lean | 10 ++ .../Localization/Away/AdjoinRoot.lean | 2 +- .../RingTheory/Localization/Away/Basic.lean | 110 +++++++++++++++--- 5 files changed, 152 insertions(+), 56 deletions(-) diff --git a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean index 87195c4877f61..188473407ba08 100644 --- a/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean +++ b/Mathlib/AlgebraicGeometry/GammaSpecAdjunction.lean @@ -136,7 +136,7 @@ theorem toΓSpecCApp_iff -- Porting Note: Type class problem got stuck in `IsLocalization.Away.AwayMap.lift_comp` -- created instance manually. This replaces the `pick_goal` tactics have loc_inst := IsLocalization.to_basicOpen (Γ.obj (op X)) r - rw [← @IsLocalization.Away.AwayMap.lift_comp _ _ _ _ _ _ _ r loc_inst _ + rw [← @IsLocalization.Away.lift_comp _ _ _ _ _ _ _ r loc_inst _ (X.isUnit_res_toΓSpecMapBasicOpen r)] --pick_goal 5; exact is_localization.to_basic_open _ r constructor diff --git a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean index 6a4899d2598b7..bb1699add44f8 100644 --- a/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean +++ b/Mathlib/AlgebraicGeometry/PrimeSpectrum/Basic.lean @@ -9,7 +9,7 @@ import Mathlib.RingTheory.Ideal.Over import Mathlib.RingTheory.KrullDimension.Basic import Mathlib.RingTheory.LocalRing.ResidueField.Defs import Mathlib.RingTheory.LocalRing.RingHom.Basic -import Mathlib.RingTheory.Localization.Away.Lemmas +import Mathlib.RingTheory.Localization.Away.Basic import Mathlib.Tactic.StacksAttribute import Mathlib.Topology.KrullDimension import Mathlib.Topology.Sober @@ -530,18 +530,61 @@ theorem isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton [Algebra R S] {f : R} {p : PrimeSpectrum R} (h : (basicOpen f).1 = {p}) : IsLocalization.Away f S ↔ IsLocalization.AtPrime S p.1 := have : IsLocalization.AtPrime (Localization.Away f) p.1 := by - refine .of_le_of_exists_dvd (Submonoid.powers f) _ + refine .of_le_of_exists_dvd (.powers f) _ (Submonoid.powers_le.mpr <| by apply h ▸ Set.mem_singleton p) fun r hr ↦ ?_ contrapose! hr simp_rw [← Ideal.mem_span_singleton] at hr have ⟨q, prime, le, disj⟩ := Ideal.exists_le_prime_disjoint (Ideal.span {r}) - (Submonoid.powers f) (Set.disjoint_right.mpr hr) + (.powers f) (Set.disjoint_right.mpr hr) have : ⟨q, prime⟩ ∈ (basicOpen f).1 := Set.disjoint_right.mp disj (Submonoid.mem_powers f) rw [h, Set.mem_singleton_iff] at this rw [← this] exact not_not.mpr (q.span_singleton_le_iff_mem.mp le) IsLocalization.isLocalization_iff_of_isLocalization _ _ (Localization.Away f) +variable [DiscreteTopology (PrimeSpectrum R)] + +variable (R) in +lemma _root_.RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : + Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by + have (p : PrimeSpectrum R) : ∃ f, (basicOpen f : Set _) = {p} := + have ⟨_, ⟨f, rfl⟩, hpf, hfp⟩ := isTopologicalBasis_basic_opens.isOpen_iff.mp + (isOpen_discrete {p}) p rfl + ⟨f, hfp.antisymm <| Set.singleton_subset_iff.mpr hpf⟩ + choose f hf using this + let e := Equiv.ofInjective f fun p q eq ↦ Set.singleton_injective (hf p ▸ eq ▸ hf q) + have loc a : IsLocalization.AtPrime (Localization.Away a.1) (e.symm a).1 := + (isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton <| hf _).mp <| by + simp_rw [Equiv.apply_ofInjective_symm]; infer_instance + let algE a := IsLocalization.algEquiv (e.symm a).1.primeCompl + (Localization.AtPrime (e.symm a).1) (Localization.Away a.1) + have span_eq : Ideal.span (Set.range f) = ⊤ := iSup_basicOpen_eq_top_iff.mp <| top_unique + fun p _ ↦ TopologicalSpace.Opens.mem_iSup.mpr ⟨p, (hf p).ge rfl⟩ + replace hf a : (basicOpen a.1 : Set _) = {e.symm a} := by + simp_rw [← hf, Equiv.apply_ofInjective_symm] + have := (discreteTopology_iff_finite_and_isPrime_imp_isMaximal.mp ‹_›).2 + obtain ⟨r, eq, -⟩ := Localization.existsUnique_algebraMap_eq_of_span_eq_top _ span_eq + (fun a ↦ algE a (x ⟨_, this _ inferInstance⟩)) fun a b ↦ by + obtain rfl | ne := eq_or_ne a b; · rfl + have ⟨n, hn⟩ : IsNilpotent (a * b : R) := (basicOpen_eq_bot_iff _).mp <| by + simp_rw [basicOpen_mul, SetLike.ext'_iff, TopologicalSpace.Opens.coe_inf, hf] + exact bot_unique (fun _ ⟨ha, hb⟩ ↦ ne <| e.symm.injective (ha.symm.trans hb)) + have := IsLocalization.subsingleton (M := .powers (a * b : R)) + (S := Localization.Away (a * b : R)) <| hn ▸ ⟨n, rfl⟩ + apply Subsingleton.elim + refine ⟨r, funext fun I ↦ ?_⟩ + have := eq (e ⟨I, I.2.isPrime⟩) + rwa [← AlgEquiv.symm_apply_eq, AlgEquiv.commutes, e.symm_apply_apply] at this + +/-- If the prime spectrum of a commutative semiring R has discrete Zariski topology, then R is +canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ +@[stacks 00JA +"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] +def _root_.RingHom.toLocalizationIsMaximalEquiv : R ≃+* + Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := + .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, + RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ + end BasicOpen section Order @@ -1102,39 +1145,4 @@ lemma isIdempotentElemEquivClopens_symm_sup (s₁ s₂ : Clopens (PrimeSpectrum end PrimeSpectrum -variable [DiscreteTopology (PrimeSpectrum R)] -open PrimeSpectrum - -variable (R) in -lemma RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology : - Function.Surjective (RingHom.toLocalizationIsMaximal R) := fun x ↦ by - let idem I := isIdempotentElemEquivClopens (R := R).symm ⟨{I}, isClopen_discrete _⟩ - let ideal I := Ideal.span {1 - (idem I).1} - let toSpec (I : {I : Ideal R | I.IsMaximal}) : PrimeSpectrum R := ⟨I.1, I.2.isPrime⟩ - have loc I : IsLocalization.AtPrime (R ⧸ ideal I) I.1 := by - rw [← isLocalization_away_iff_atPrime_of_basicOpen_eq_singleton] - exacts [IsLocalization.Away.quotient_of_isIdempotentElem (idem I).2, - congr_arg (·.1) (basicOpen_isIdempotentElemEquivClopens_symm ⟨{I}, isClopen_discrete _⟩)] - let equiv I := IsLocalization.algEquiv I.1.primeCompl (Localization.AtPrime I.1) (R ⧸ ideal I) - have := (discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical.mp ‹_›).1 - have ⟨r, hr⟩ := Ideal.pi_quotient_surjective ?_ fun I ↦ equiv (toSpec I) (x I) - · refine ⟨r, funext fun I ↦ (equiv <| toSpec I).injective ?_⟩ - rw [← hr]; exact (equiv _).commutes r - refine fun I J ne ↦ Ideal.isCoprime_iff_exists.mpr ?_ - have := ((idem <| toSpec I).2.mul (idem <| toSpec J).2).eq_zero_of_isNilpotent <| by - simp_rw [← basicOpen_eq_bot_iff, basicOpen_mul, SetLike.ext'_iff, idem, - TopologicalSpace.Opens.coe_inf, basicOpen_isIdempotentElemEquivClopens_symm] - exact Set.singleton_inter_eq_empty.mpr fun h ↦ ne (Subtype.ext <| congr_arg (·.1) h) - simp_rw [ideal, Ideal.mem_span_singleton', exists_exists_eq_and] - exact ⟨1, idem (toSpec I), by simpa [mul_sub]⟩ - -/-- If the prime spectrum of a commutative ring R has discrete Zariski topology, then R is -canonically isomorphic to the product of its localizations at the (finitely many) maximal ideals. -/ -@[stacks 00JA -"See also `PrimeSpectrum.discreteTopology_iff_finite_isMaximal_and_sInf_le_nilradical`."] -def RingHom.toLocalizationIsMaximalEquiv : R ≃+* - Π I : {I : Ideal R // I.IsMaximal}, haveI : I.1.IsMaximal := I.2; Localization.AtPrime I.1 := - .ofBijective _ ⟨RingHom.toLocalizationIsMaximal_injective R, - RingHom.toLocalizationIsMaximal_surjective_of_discreteTopology R⟩ - end Idempotent diff --git a/Mathlib/RingTheory/Ideal/Basic.lean b/Mathlib/RingTheory/Ideal/Basic.lean index 0c7c7a913ae1b..9bcac7b686908 100644 --- a/Mathlib/RingTheory/Ideal/Basic.lean +++ b/Mathlib/RingTheory/Ideal/Basic.lean @@ -150,6 +150,16 @@ theorem span_pow_eq_top (s : Set α) (hs : span s = ⊤) (n : ℕ) : rw [mul_pow, mem_span_singleton] exact ⟨f x ^ (n + 1), mul_comm _ _⟩ +theorem span_range_pow_eq_top (s : Set α) (hs : span s = ⊤) (n : s → ℕ) : + span (Set.range fun x ↦ x.1 ^ n x) = ⊤ := by + have ⟨t, hts, mem⟩ := Submodule.mem_span_finite_of_mem_span ((eq_top_iff_one _).mp hs) + refine top_unique ((span_pow_eq_top _ ((eq_top_iff_one _).mpr mem) <| + t.attach.sup fun x ↦ n ⟨x, hts x.2⟩).ge.trans <| span_le.mpr ?_) + rintro _ ⟨x, hxt, rfl⟩ + rw [← Nat.sub_add_cancel (Finset.le_sup <| t.mem_attach ⟨x, hxt⟩)] + simp_rw [pow_add] + exact mul_mem_left _ _ (subset_span ⟨_, rfl⟩) + end Ideal end CommSemiring diff --git a/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean b/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean index 5211fab000937..4061c6c8231de 100644 --- a/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean +++ b/Mathlib/RingTheory/Localization/Away/AdjoinRoot.lean @@ -28,7 +28,7 @@ noncomputable def Localization.awayEquivAdjoin (r : R) : Away r ≃ₐ[R] Adjoin (isUnit_of_mul_eq_one ((algebraMap R (AdjoinRoot (C r * X - 1))) r) (root (C r * X - 1)) (root_isInv r)) with commutes' := - IsLocalization.Away.AwayMap.lift_eq r (isUnit_of_mul_eq_one _ _ <| root_isInv r) } + IsLocalization.Away.lift_eq r (isUnit_of_mul_eq_one _ _ <| root_isInv r) } (liftHom _ (IsLocalization.Away.invSelf r) <| by simp only [map_sub, map_mul, aeval_C, aeval_X, IsLocalization.Away.mul_invSelf, aeval_one, sub_self]) diff --git a/Mathlib/RingTheory/Localization/Away/Basic.lean b/Mathlib/RingTheory/Localization/Away/Basic.lean index ae069a4b51072..259a7cba398d2 100644 --- a/Mathlib/RingTheory/Localization/Away/Basic.lean +++ b/Mathlib/RingTheory/Localization/Away/Basic.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau, Mario Carneiro, Johan Commelin, Amelia Livingston, Anne Baanen -/ import Mathlib.GroupTheory.MonoidLocalization.Away +import Mathlib.Algebra.Algebra.Pi import Mathlib.RingTheory.Ideal.Maps import Mathlib.RingTheory.Localization.Basic import Mathlib.RingTheory.UniqueFactorizationDomain.Multiplicity @@ -123,6 +124,9 @@ lemma iff_of_associated {r r' : R} (h : Associated r r') : IsLocalization.Away r S ↔ IsLocalization.Away r' S := ⟨fun _ ↦ IsLocalization.Away.of_associated h, fun _ ↦ IsLocalization.Away.of_associated h.symm⟩ +lemma isUnit_of_dvd {r : R} (h : r ∣ x) : IsUnit (algebraMap R S r) := + isUnit_of_dvd_unit (map_dvd _ h) (algebraMap_isUnit x) + variable {g : R →+* P} /-- Given `x : R`, a localization map `F : R →+* S` away from `x`, and a map of `CommSemiring`s @@ -136,20 +140,33 @@ noncomputable def lift (hg : IsUnit (g x)) : S →+* P := exact IsUnit.map (powMonoidHom n : P →* P) hg @[simp] -theorem AwayMap.lift_eq (hg : IsUnit (g x)) (a : R) : lift x hg ((algebraMap R S) a) = g a := +theorem lift_eq (hg : IsUnit (g x)) (a : R) : lift x hg (algebraMap R S a) = g a := IsLocalization.lift_eq _ _ @[simp] -theorem AwayMap.lift_comp (hg : IsUnit (g x)) : (lift x hg).comp (algebraMap R S) = g := +theorem lift_comp (hg : IsUnit (g x)) : (lift x hg).comp (algebraMap R S) = g := IsLocalization.lift_comp _ +@[deprecated (since := "2024-11-25")] alias AwayMap.lift_eq := lift_eq +@[deprecated (since := "2024-11-25")] alias AwayMap.lift_comp := lift_comp + +/-- Given `x y : R` and localizations `S`, `P` away from `x` and `y * x` +respectively, the homomorphism induced from `S` to `P`. -/ +noncomputable def awayToAwayLeft (y : R) [Algebra R P] [IsLocalization.Away (y * x) P] : S →+* P := + lift x <| isUnit_of_dvd (y * x) (dvd_mul_left _ _) + /-- Given `x y : R` and localizations `S`, `P` away from `x` and `x * y` respectively, the homomorphism induced from `S` to `P`. -/ noncomputable def awayToAwayRight (y : R) [Algebra R P] [IsLocalization.Away (x * y) P] : S →+* P := - lift x <| - show IsUnit ((algebraMap R P) x) from - isUnit_of_mul_eq_one ((algebraMap R P) x) (mk' P y ⟨x * y, Submonoid.mem_powers _⟩) <| by - rw [mul_mk'_eq_mk'_of_mul, mk'_self] + lift x <| isUnit_of_dvd (x * y) (dvd_mul_right _ _) + +theorem awayToAwayLeft_eq (y : R) [Algebra R P] [IsLocalization.Away (y * x) P] (a : R) : + awayToAwayLeft x y (algebraMap R S a) = algebraMap R P a := + lift_eq _ _ _ + +theorem awayToAwayRight_eq (y : R) [Algebra R P] [IsLocalization.Away (x * y) P] (a : R) : + awayToAwayRight x y (algebraMap R S a) = algebraMap R P a := + lift_eq _ _ _ variable (S) (Q : Type*) [CommSemiring Q] [Algebra P Q] @@ -164,10 +181,10 @@ noncomputable def map (f : R →+* P) (r : R) [IsLocalization.Away r S] section Algebra -variable {A : Type*} [CommRing A] [Algebra R A] -variable {B : Type*} [CommRing B] [Algebra R B] -variable (Aₚ : Type*) [CommRing Aₚ] [Algebra A Aₚ] [Algebra R Aₚ] [IsScalarTower R A Aₚ] -variable (Bₚ : Type*) [CommRing Bₚ] [Algebra B Bₚ] [Algebra R Bₚ] [IsScalarTower R B Bₚ] +variable {A : Type*} [CommSemiring A] [Algebra R A] +variable {B : Type*} [CommSemiring B] [Algebra R B] +variable (Aₚ : Type*) [CommSemiring Aₚ] [Algebra A Aₚ] [Algebra R Aₚ] [IsScalarTower R A Aₚ] +variable (Bₚ : Type*) [CommSemiring Bₚ] [Algebra B Bₚ] [Algebra R Bₚ] [IsScalarTower R B Bₚ] instance {f : A →+* B} (a : A) [Away (f a) Bₚ] : IsLocalization (.map f (.powers a)) Bₚ := by simpa @@ -279,7 +296,7 @@ noncomputable def atUnit (x : R) (e : IsUnit x) [IsLocalization.Away x S] : R noncomputable def atOne [IsLocalization.Away (1 : R) S] : R ≃ₐ[R] S := @atUnit R _ S _ _ (1 : R) isUnit_one _ -theorem away_of_isUnit_of_bijective {R : Type*} (S : Type*) [CommRing R] [CommRing S] +theorem away_of_isUnit_of_bijective {R : Type*} (S : Type*) [CommSemiring R] [CommSemiring S] [Algebra R S] {r : R} (hr : IsUnit r) (H : Function.Bijective (algebraMap R S)) : IsLocalization.Away r S := { map_units' := by @@ -364,28 +381,88 @@ noncomputable abbrev awayMap (f : R →+* P) (r : R) : Localization.Away r →+* Localization.Away (f r) := IsLocalization.Away.map _ _ f r -variable {A : Type*} [CommRing A] [Algebra R A] -variable {B : Type*} [CommRing B] [Algebra R B] +variable {A : Type*} [CommSemiring A] [Algebra R A] +variable {B : Type*} [CommSemiring B] [Algebra R B] /-- Given a map `f : A →ₐ[R] B` and an element `a : A`, we may construct a map `Aₐ →ₐ[R] Bₐ`. -/ noncomputable abbrev awayMapₐ (f : A →ₐ[R] B) (a : A) : Localization.Away a →ₐ[R] Localization.Away (f a) := IsLocalization.Away.mapₐ _ _ f a +theorem algebraMap_injective_of_span_eq_top (s : Set R) (span_eq : Ideal.span s = ⊤) : + Function.Injective (algebraMap R <| Π a : s, Away a.1) := fun x y eq ↦ by + suffices Module.eqIdeal R x y = ⊤ by simpa [Module.eqIdeal] using (Ideal.eq_top_iff_one _).mp this + by_contra ne + have ⟨r, hrs, disj⟩ := Ideal.exists_disjoint_powers_of_span_eq_top s span_eq _ ne + let r : s := ⟨r, hrs⟩ + have ⟨⟨_, n, rfl⟩, eq⟩ := (IsLocalization.eq_iff_exists (.powers r.1) _).mp (congr_fun eq r) + exact Set.disjoint_left.mp disj eq ⟨n, rfl⟩ + +/-- The sheaf condition for the structure sheaf on `Spec R` +for a covering of the whole prime spectrum by basic opens. -/ +theorem existsUnique_algebraMap_eq_of_span_eq_top (s : Set R) (span_eq : Ideal.span s = ⊤) + (f : Π a : s, Away a.1) (h : ∀ a b : s, + Away.awayToAwayRight (P := Away (a * b : R)) a.1 b (f a) = Away.awayToAwayLeft b.1 a (f b)) : + ∃! r : R, ∀ a : s, algebraMap R _ r = f a := by + have mem := (Ideal.eq_top_iff_one _).mp span_eq; clear span_eq + wlog finset_eq : ∃ t : Finset R, t = s generalizing s + · have ⟨t, hts, mem⟩ := Submodule.mem_span_finite_of_mem_span mem + have ⟨r, eq, uniq⟩ := this t (fun a ↦ f ⟨a, hts a.2⟩) + (fun a b ↦ h ⟨a, hts a.2⟩ ⟨b, hts b.2⟩) mem ⟨_, rfl⟩ + refine ⟨r, fun a ↦ ?_, fun _ eq ↦ uniq _ fun a ↦ eq ⟨a, hts a.2⟩⟩ + replace hts := Set.insert_subset a.2 hts + classical + have ⟨r', eq, _⟩ := this ({a.1} ∪ t) (fun a ↦ f ⟨a, hts a.2⟩) (fun a b ↦ + h ⟨a, hts a.2⟩ ⟨b, hts b.2⟩) (Ideal.span_mono (fun _ ↦ .inr) mem) ⟨{a.1} ∪ t, by simp⟩ + exact (congr_arg _ (uniq _ fun b ↦ eq ⟨b, .inr b.2⟩).symm).trans (eq ⟨a, .inl rfl⟩) + have span_eq := (Ideal.eq_top_iff_one _).mpr mem + refine exists_unique_of_exists_of_unique ?_ fun x y hx hy ↦ + algebraMap_injective_of_span_eq_top s span_eq (funext fun a ↦ (hx a).trans (hy a).symm) + obtain ⟨s, rfl⟩ := finset_eq + choose n r eq using fun a ↦ Away.surj a.1 (f a) + let N := s.attach.sup n + let r a := a ^ (N - n a) * r a + have eq a : f a * algebraMap R _ (a ^ N) = algebraMap R _ (r a) := by + rw [map_mul, ← eq, mul_left_comm, ← map_pow, ← map_mul, ← pow_add, + Nat.sub_add_cancel (Finset.le_sup <| s.mem_attach a)] + have eq2 a b : ∃ N', (a * b) ^ N' * (r a * b ^ N) = (a * b) ^ N' * (r b * a ^ N) := + Away.exists_of_eq (S := Away (a * b : R)) _ <| by + simp_rw [map_mul, ← Away.awayToAwayRight_eq (S := Away a.1) a.1 b (r a), + ← Away.awayToAwayLeft_eq (S := Away b.1) b.1 a (r b), ← eq, map_mul, + Away.awayToAwayRight_eq, Away.awayToAwayLeft_eq, h, mul_assoc, ← map_mul, mul_comm] + choose N' hN' using eq2 + let N' := (s ×ˢ s).attach.sup fun a ↦ N' + ⟨_, (Finset.mem_product.mp a.2).1⟩ ⟨_, (Finset.mem_product.mp a.2).2⟩ + have eq2 a b : (a * b) ^ N' * (r a * b ^ N) = (a * b) ^ N' * (r b * a ^ N) := by + dsimp only [N']; rw [← Nat.sub_add_cancel (Finset.le_sup <| (Finset.mem_attach _ ⟨⟨a, b⟩, + Finset.mk_mem_product a.2 b.2⟩)), pow_add, mul_assoc, hN', ← mul_assoc] + let N := N' + N + let r a := a ^ N' * r a + have eq a : f a * algebraMap R _ (a ^ N) = algebraMap R _ (r a) := by + rw [map_mul, ← eq, mul_left_comm, ← map_mul, ← pow_add] + have eq2 a b : r a * b ^ N = r b * a ^ N := by + rw [pow_add, mul_mul_mul_comm, ← mul_pow, eq2, + mul_comm a.1, mul_pow, mul_mul_mul_comm, ← pow_add] + have ⟨c, eq1⟩ := (mem_span_range_iff_exists_fun _).mp <| + (Ideal.eq_top_iff_one _).mp <| (Set.image_eq_range _ _ ▸ Ideal.span_pow_eq_top _ span_eq N) + refine ⟨∑ b, c b * r b, fun a ↦ ((Away.algebraMap_isUnit a.1).pow N).mul_left_inj.mp ?_⟩ + simp_rw [← map_pow, eq, ← map_mul, Finset.sum_mul, mul_assoc, eq2 _ a, mul_left_comm (c _), + ← Finset.mul_sum, ← smul_eq_mul (a := c _), eq1, mul_one] + end Localization end CommSemiring open Localization -variable {R : Type*} [CommRing R] +variable {R : Type*} [CommSemiring R] section NumDen open IsLocalization variable (x : R) -variable (B : Type*) [CommRing B] [Algebra R B] [IsLocalization.Away x B] +variable (B : Type*) [CommSemiring B] [Algebra R B] [IsLocalization.Away x B] /-- `selfZPow x (m : ℤ)` is `x ^ m` as an element of the localization away from `x`. -/ noncomputable def selfZPow (m : ℤ) : B := @@ -478,7 +555,8 @@ theorem selfZPow_pow_sub (a : R) (b : B) (m d : ℤ) : simp only at this rwa [mul_comm _ b, mul_assoc b _ _, selfZPow_mul_neg, mul_one] at this -variable [IsDomain R] [WfDvdMonoid R] +variable {R : Type*} [CommRing R] (x : R) (B : Type*) [CommRing B] +variable [Algebra R B] [IsLocalization.Away x B] [IsDomain R] [WfDvdMonoid R] theorem exists_reduced_fraction' {b : B} (hb : b ≠ 0) (hx : Irreducible x) : ∃ (a : R) (n : ℤ), ¬x ∣ a ∧ selfZPow x B n * algebraMap R B a = b := by From de39dd8ecfea9b240dede98cdbbedf44f906a37f Mon Sep 17 00:00:00 2001 From: damiano Date: Tue, 26 Nov 2024 22:40:39 +0000 Subject: [PATCH 201/250] feat: zulip emoji reactions on public and `mathlib reviewers` channels (#19522) Updates the action that places emoji reactions triggered by `bors x`: now the reactions should be added to all public channels and to the `mathlib reviewers` channel. --- scripts/zulip_emoji_merge_delegate.py | 29 ++++++++++++++++++++++----- 1 file changed, 24 insertions(+), 5 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index e572aeb90e8d3..8fa4e33e74f4c 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,15 +23,34 @@ site=ZULIP_SITE ) -# Fetch the last 200 messages -response = client.get_messages({ +# Fetch the messages containing the PR number from the public channels. +# There does not seem to be a way to search simultaneously public and private channels. +public_response = client.get_messages({ "anchor": "newest", - "num_before": 200, + "num_before": 5000, "num_after": 0, - "narrow": [{"operator": "channel", "operand": "PR reviews"}], + "narrow": [ + {"operator": "channels", "operand": "public"}, + {"operator": "search", "operand": f'#{PR_NUMBER}'}, + ], }) -messages = response['messages'] +# Fetch the messages containing the PR number from the `mathlib reviewers` channel +# There does not seem to be a way to search simultaneously public and private channels. +reviewers_response = client.get_messages({ + "anchor": "newest", + "num_before": 5000, + "num_after": 0, + "narrow": [ + {"operator": "channel", "operand": "mathlib reviewers"}, + {"operator": "search", "operand": f'#{PR_NUMBER}'}, + ], +}) + +print(f"public_response:{public_response}") +print(f"reviewers_response:{reviewers_response}") + +messages = (public_response['messages']) + (reviewers_response['messages']) pr_pattern = re.compile(f'https://github.com/leanprover-community/mathlib4/pull/{PR_NUMBER}') From 84b098419d680a531472f202a5ce5234bec118b6 Mon Sep 17 00:00:00 2001 From: peabrainiac Date: Tue, 26 Nov 2024 23:03:57 +0000 Subject: [PATCH 202/250] feat(Topology/Compactness): delta-generated spaces (#19431) Introduces delta-generated topological spaces and shows that they are sequential, locally path-connected and closed under disjoint unions and quotients. Co-authored-by: peabrainiac <43812953+peabrainiac@users.noreply.github.com> --- Mathlib.lean | 1 + .../Compactness/DeltaGeneratedSpace.lean | 183 ++++++++++++++++++ 2 files changed, 184 insertions(+) create mode 100644 Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean diff --git a/Mathlib.lean b/Mathlib.lean index c971dda66ba49..1fe86e9134e3b 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -5033,6 +5033,7 @@ import Mathlib.Topology.Compactification.OnePoint import Mathlib.Topology.Compactification.OnePointEquiv import Mathlib.Topology.Compactness.Compact import Mathlib.Topology.Compactness.CompactlyGeneratedSpace +import Mathlib.Topology.Compactness.DeltaGeneratedSpace import Mathlib.Topology.Compactness.Exterior import Mathlib.Topology.Compactness.Lindelof import Mathlib.Topology.Compactness.LocallyCompact diff --git a/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean new file mode 100644 index 0000000000000..c7c9e34c6a6ae --- /dev/null +++ b/Mathlib/Topology/Compactness/DeltaGeneratedSpace.lean @@ -0,0 +1,183 @@ +/- +Copyright (c) 2024 Ben Eltschig. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Ben Eltschig +-/ +import Mathlib.Analysis.Convex.Normed + +/-! +# Delta-generated topological spaces + +This file defines delta-generated spaces, as topological spaces whose topology is coinduced by all +maps from euclidean spaces into them. This is the strongest topological property that holds for +all CW-complexes and is closed under quotients and disjoint unions; every delta-generated space is +locally path-connected, sequential and in particular compactly generated. + +See https://ncatlab.org/nlab/show/Delta-generated+topological+space. + +Adapted from `Mathlib.Topology.Compactness.CompactlyGeneratedSpace`. + +## TODO +* All locally path-connected first-countable spaces are delta-generated - in particular, all normed + spaces and convex subsets thereof, like the standard simplices and the unit interval. +* Delta-generated spaces are equivalently generated by the simplices Δⁿ. +* Delta-generated spaces are equivalently generated by the unit interval I. +-/ + +variable {X Y : Type*} [tX : TopologicalSpace X] [tY : TopologicalSpace Y] + +open TopologicalSpace Topology + +/-- The topology coinduced by all maps from ℝⁿ into a space. -/ +def TopologicalSpace.deltaGenerated (X : Type*) [TopologicalSpace X] : TopologicalSpace X := + ⨆ f : (n : ℕ) × C(((Fin n) → ℝ), X), coinduced f.2 inferInstance + +/-- The delta-generated topology is also coinduced by a single map out of a sigma type. -/ +lemma deltaGenerated_eq_coinduced : deltaGenerated X = coinduced + (fun x : (f : (n : ℕ) × C(Fin n → ℝ, X)) × (Fin f.1 → ℝ) ↦ x.1.2 x.2) inferInstance := by + rw [deltaGenerated, instTopologicalSpaceSigma, coinduced_iSup]; rfl + +/-- The delta-generated topology is at least as fine as the original one. -/ +lemma deltaGenerated_le : deltaGenerated X ≤ tX := + iSup_le_iff.mpr fun f ↦ f.2.continuous.coinduced_le + +/-- A set is open in `deltaGenerated X` iff all its preimages under continuous functions ℝⁿ → X are + open. -/ +lemma isOpen_deltaGenerated_iff {u : Set X} : + IsOpen[deltaGenerated X] u ↔ ∀ n (p : C(Fin n → ℝ, X)), IsOpen (p ⁻¹' u) := by + simp_rw [deltaGenerated, isOpen_iSup_iff, isOpen_coinduced, Sigma.forall] + +/-- A map from ℝⁿ to X is continuous iff it is continuous regarding the + delta-generated topology on X. Outside of this file, use the more general + `continuous_to_deltaGenerated` instead. -/ +private lemma continuous_euclidean_to_deltaGenerated {n : ℕ} {f : (Fin n → ℝ) → X} : + Continuous[_, deltaGenerated X] f ↔ Continuous f := by + simp_rw [continuous_iff_coinduced_le] + refine ⟨fun h ↦ h.trans deltaGenerated_le, fun h ↦ ?_⟩ + simp_rw [deltaGenerated] + exact le_iSup_of_le (i := ⟨n, f, continuous_iff_coinduced_le.mpr h⟩) le_rfl + +/-- `deltaGenerated` is idempotent as a function `TopologicalSpace X → TopologicalSpace X`. -/ +lemma deltaGenerated_deltaGenerated_eq : + @deltaGenerated X (deltaGenerated X) = deltaGenerated X := by + ext u; simp_rw [isOpen_deltaGenerated_iff]; refine forall_congr' fun n ↦ ?_ + -- somewhat awkward because `ContinuousMap` doesn't play well with multiple topologies. + refine ⟨fun h p ↦ h <| @ContinuousMap.mk _ _ _ (_) p ?_, fun h p ↦ h ⟨p, ?_⟩⟩ + · exact continuous_euclidean_to_deltaGenerated.mpr p.2 + · exact continuous_euclidean_to_deltaGenerated.mp <| @ContinuousMap.continuous_toFun _ _ _ (_) p + +/-- A space is delta-generated if its topology is equal to the delta-generated topology, i.e. + coinduced by all continuous maps ℝⁿ → X. Since the delta-generated topology is always finer + than the original one, it suffices to show that it is also coarser. -/ +class DeltaGeneratedSpace (X : Type*) [t : TopologicalSpace X] : Prop where + le_deltaGenerated : t ≤ deltaGenerated X + +lemma eq_deltaGenerated [DeltaGeneratedSpace X] : tX = deltaGenerated X := + eq_of_le_of_le DeltaGeneratedSpace.le_deltaGenerated deltaGenerated_le + +/-- A subset of a delta-generated space is open iff its preimage is open for every + continuous map from ℝⁿ to X. -/ +lemma DeltaGeneratedSpace.isOpen_iff [DeltaGeneratedSpace X] {u : Set X} : + IsOpen u ↔ ∀ (n : ℕ) (p : ContinuousMap ((Fin n) → ℝ) X), IsOpen (p ⁻¹' u) := by + nth_rewrite 1 [eq_deltaGenerated (X := X)]; exact isOpen_deltaGenerated_iff + +/-- A map out of a delta-generated space is continuous iff it preserves continuity of maps + from ℝⁿ into X. -/ +lemma DeltaGeneratedSpace.continuous_iff [DeltaGeneratedSpace X] {f : X → Y} : + Continuous f ↔ ∀ (n : ℕ) (p : C(((Fin n) → ℝ), X)), Continuous (f ∘ p) := by + simp_rw [continuous_iff_coinduced_le] + nth_rewrite 1 [eq_deltaGenerated (X := X), deltaGenerated] + simp [coinduced_compose] + constructor + · intro h n p; apply h ⟨n, p⟩ + · rintro h ⟨n, p⟩; apply h n p + +/-- A map out of a delta-generated space is continuous iff it is continuous with respect + to the delta-generated topology on the codomain. -/ +lemma continuous_to_deltaGenerated [DeltaGeneratedSpace X] {f : X → Y} : + Continuous[_, deltaGenerated Y] f ↔ Continuous f := by + simp_rw [DeltaGeneratedSpace.continuous_iff, continuous_euclidean_to_deltaGenerated] + +/-- The delta-generated topology on `X` does in fact turn `X` into a delta-generated space. -/ +lemma deltaGeneratedSpace_deltaGenerated {X : Type*} {t : TopologicalSpace X} : + @DeltaGeneratedSpace X (@deltaGenerated X t) := by + let _ := @deltaGenerated X t; constructor; rw [@deltaGenerated_deltaGenerated_eq X t] + +lemma deltaGenerated_mono {X : Type*} {t₁ t₂ : TopologicalSpace X} (h : t₁ ≤ t₂) : + @deltaGenerated X t₁ ≤ @deltaGenerated X t₂ := by + rw [← continuous_id_iff_le, @continuous_to_deltaGenerated _ _ + (@deltaGenerated X t₁) t₂ deltaGeneratedSpace_deltaGenerated id] + exact continuous_id_iff_le.2 <| (@deltaGenerated_le X t₁).trans h + +namespace DeltaGeneratedSpace + +/-- Type synonym to be equipped with the delta-generated topology. -/ +def of (X : Type*) := X + +instance : TopologicalSpace (of X) := deltaGenerated X + +instance : DeltaGeneratedSpace (of X) := + deltaGeneratedSpace_deltaGenerated + +/-- The natural map from `DeltaGeneratedSpace.of X` to `X`. -/ +def counit : (of X) → X := id + +lemma continuous_counit : Continuous (counit : _ → X) := by + rw [continuous_iff_coinduced_le]; exact deltaGenerated_le + +/-- Delta-generated spaces are locally path-connected. -/ +instance [DeltaGeneratedSpace X] : LocPathConnectedSpace X := by + rw [eq_deltaGenerated (X := X), deltaGenerated_eq_coinduced] + exact LocPathConnectedSpace.coinduced _ + +/-- Delta-generated spaces are sequential. -/ +instance [DeltaGeneratedSpace X] : SequentialSpace X := by + rw [eq_deltaGenerated (X := X)] + exact SequentialSpace.iSup fun p ↦ SequentialSpace.coinduced p.2 + +end DeltaGeneratedSpace + +omit tY in +/-- Any topology coinduced by a delta-generated topology is delta-generated. -/ +lemma DeltaGeneratedSpace.coinduced [DeltaGeneratedSpace X] (f : X → Y) : + @DeltaGeneratedSpace Y (tX.coinduced f) := + let _ := tX.coinduced f + ⟨(continuous_to_deltaGenerated.2 continuous_coinduced_rng).coinduced_le⟩ + +/-- Suprema of delta-generated topologies are delta-generated. -/ +protected lemma DeltaGeneratedSpace.iSup {X : Type*} {ι : Sort*} {t : ι → TopologicalSpace X} + (h : ∀ i, @DeltaGeneratedSpace X (t i)) : @DeltaGeneratedSpace X (⨆ i, t i) := + let _ := ⨆ i, t i + ⟨iSup_le_iff.2 fun i ↦ (h i).le_deltaGenerated.trans <| deltaGenerated_mono <| le_iSup t i⟩ + +/-- Suprema of delta-generated topologies are delta-generated. -/ +protected lemma DeltaGeneratedSpace.sup {X : Type*} {t₁ t₂ : TopologicalSpace X} + (h₁ : @DeltaGeneratedSpace X t₁) (h₂ : @DeltaGeneratedSpace X t₂) : + @DeltaGeneratedSpace X (t₁ ⊔ t₂) := by + rw [sup_eq_iSup] + exact .iSup <| Bool.forall_bool.2 ⟨h₂, h₁⟩ + +/-- Quotients of delta-generated spaces are delta-generated. -/ +lemma Topology.IsQuotientMap.deltaGeneratedSpace [DeltaGeneratedSpace X] + {f : X → Y} (h : IsQuotientMap f) : DeltaGeneratedSpace Y := + h.2 ▸ DeltaGeneratedSpace.coinduced f + +/-- Quotients of delta-generated spaces are delta-generated. -/ +instance Quot.deltaGeneratedSpace [DeltaGeneratedSpace X] {r : X → X → Prop} : + DeltaGeneratedSpace (Quot r) := + isQuotientMap_quot_mk.deltaGeneratedSpace + +/-- Quotients of delta-generated spaces are delta-generated. -/ +instance Quotient.deltaGeneratedSpace [DeltaGeneratedSpace X] {s : Setoid X} : + DeltaGeneratedSpace (Quotient s) := + isQuotientMap_quotient_mk'.deltaGeneratedSpace + +/-- Disjoint unions of delta-generated spaces are delta-generated. -/ +instance Sum.deltaGeneratedSpace [DeltaGeneratedSpace X] [DeltaGeneratedSpace Y] : + DeltaGeneratedSpace (X ⊕ Y) := + DeltaGeneratedSpace.sup (.coinduced Sum.inl) (.coinduced Sum.inr) + +/-- Disjoint unions of delta-generated spaces are delta-generated. -/ +instance Sigma.deltaGeneratedSpace {ι : Type*} {X : ι → Type*} [∀ i, TopologicalSpace (X i)] + [∀ i, DeltaGeneratedSpace (X i)] : DeltaGeneratedSpace (Σ i, X i) := + .iSup fun _ ↦ .coinduced _ From 1edbd1bb1d3f1789b100edb79e95ab7cc4f255f0 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 26 Nov 2024 23:03:58 +0000 Subject: [PATCH 203/250] feat(Algebra/GradedMonoid): add `prod_mem_graded` (#19500) Co-authored-by: Kevin Buzzard This contribution was created as part of the Durham Computational Algebraic Geometry Workshop. --- Mathlib/Algebra/GradedMonoid.lean | 30 ++++++++++++++++++++++++++++++ 1 file changed, 30 insertions(+) diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index 4371d6d89426e..a4aa2463b4229 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -9,6 +9,7 @@ import Mathlib.Algebra.Group.Submonoid.Defs import Mathlib.Data.List.FinRange import Mathlib.Data.SetLike.Basic import Mathlib.Data.Sigma.Basic +import Mathlib.Algebra.BigOperators.Group.Finset import Lean.Elab.Tactic /-! @@ -671,3 +672,32 @@ def SetLike.homogeneousSubmonoid [AddMonoid ι] [Monoid R] (A : ι → S) [SetLi mul_mem' a b := SetLike.homogeneous_mul a b end HomogeneousElements + +section CommMonoid + +namespace SetLike + +variable {ι R S : Type*} [SetLike S R] [CommMonoid R] [AddCommMonoid ι] +variable (A : ι → S) [SetLike.GradedMonoid A] + +variable {κ : Type*} (i : κ → ι) (g : κ → R) {F : Finset κ} + +theorem prod_mem_graded (hF : ∀ k ∈ F, g k ∈ A (i k)) : ∏ k ∈ F, g k ∈ A (∑ k ∈ F, i k) := by + classical + induction F using Finset.induction_on + · simp [GradedOne.one_mem] + · case insert j F' hF2 h3 => + rw [Finset.prod_insert hF2, Finset.sum_insert hF2] + apply SetLike.mul_mem_graded (hF j <| Finset.mem_insert_self j F') + apply h3 + intro k hk + apply hF k + exact Finset.mem_insert_of_mem hk + +theorem prod_pow_mem_graded (n : κ → ℕ) (hF : ∀ k ∈ F, g k ∈ A (i k)) : + ∏ k ∈ F, g k ^ n k ∈ A (∑ k ∈ F, n k • i k) := + prod_mem_graded A _ _ fun k hk ↦ pow_mem_graded _ (hF k hk) + +end SetLike + +end CommMonoid From ba018bd9cc58c0eff74c88301e2447b01accb47e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Tue, 26 Nov 2024 23:14:43 +0000 Subject: [PATCH 204/250] refactor(CategoryTheory/Sites): remove SheafOfTypes (#19510) This PR removes the category `SheafOfTypes.{u} J` which was equivalent to `Sheaf J (Type u)`, where `J` is a Grothendieck topology. This avoids some duplication of code, as `SheafOfTypes` was essentially a special case of `Sheaf`: the only difference was that the sheaf condition was phrased differently. --- Mathlib/CategoryTheory/Sites/Adjunction.lean | 41 +++-------- .../Sites/Coherent/LocallySurjective.lean | 15 ++-- Mathlib/CategoryTheory/Sites/Continuous.lean | 24 ++++--- .../CategoryTheory/Sites/CoverPreserving.lean | 11 +-- .../CategoryTheory/Sites/DenseSubsite.lean | 24 ++++--- .../Sites/LocallyFullyFaithful.lean | 10 +-- Mathlib/CategoryTheory/Sites/Sheaf.lean | 38 +++++----- .../CategoryTheory/Sites/SheafOfTypes.lean | 70 ------------------- Mathlib/CategoryTheory/Sites/Types.lean | 52 +++++++++----- 9 files changed, 116 insertions(+), 169 deletions(-) diff --git a/Mathlib/CategoryTheory/Sites/Adjunction.lean b/Mathlib/CategoryTheory/Sites/Adjunction.lean index 01604a3fc6ead..8b4192d40635c 100644 --- a/Mathlib/CategoryTheory/Sites/Adjunction.lean +++ b/Mathlib/CategoryTheory/Sites/Adjunction.lean @@ -29,8 +29,8 @@ variable {F : D ⥤ E} {G : E ⥤ D} /-- The forgetful functor from `Sheaf J D` to sheaves of types, for a concrete category `D` whose forgetful functor preserves the correct limits. -/ abbrev sheafForget [ConcreteCategory D] [HasSheafCompose J (forget D)] : - Sheaf J D ⥤ SheafOfTypes J := - sheafCompose J (forget D) ⋙ (sheafEquivSheafOfTypes J).functor + Sheaf J D ⥤ Sheaf J (Type _) := + sheafCompose J (forget D) namespace Sheaf @@ -98,37 +98,16 @@ section ForgetToType variable [HasWeakSheafify J D] [ConcreteCategory D] [HasSheafCompose J (forget D)] -/-- This is the functor sending a sheaf of types `X` to the sheafification of `X ⋙ G`. -/ -abbrev composeAndSheafifyFromTypes (G : Type max v₁ u₁ ⥤ D) : SheafOfTypes J ⥤ Sheaf J D := - (sheafEquivSheafOfTypes J).inverse ⋙ composeAndSheafify _ G +@[deprecated (since := "2024-11-26")] alias composeAndSheafifyFromTypes := composeAndSheafify -/-- A variant of the adjunction between sheaf categories, in the case where the right adjoint -is the forgetful functor to sheaves of types. -/ -def adjunctionToTypes {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : - composeAndSheafifyFromTypes J G ⊣ sheafForget J := - (sheafEquivSheafOfTypes J).symm.toAdjunction.comp (adjunction J adj) +/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. (Use `Sheaf.adjunction`.)-/ +@[deprecated (since := "2024-11-26")] abbrev adjunctionToTypes + {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : + composeAndSheafify J G ⊣ sheafForget J := + adjunction _ adj -@[simp] -theorem adjunctionToTypes_unit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) - (Y : SheafOfTypes J) : - ((adjunctionToTypes J adj).unit.app Y).val = - (adj.whiskerRight _).unit.app ((sheafOfTypesToPresheaf J).obj Y) ≫ - whiskerRight (toSheafify J _) (forget D) := by - simp [adjunctionToTypes] - rfl - -@[simp] -theorem adjunctionToTypes_counit_app_val {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) - (X : Sheaf J D) : - ((adjunctionToTypes J adj).counit.app X).val = - sheafifyLift J ((Functor.associator _ _ _).hom ≫ (adj.whiskerRight _).counit.app _) X.2 := by - apply sheafifyLift_unique - ext - simp [adjunctionToTypes, sheafEquivSheafOfTypes, Equivalence.symm] - -instance [(forget D).IsRightAdjoint] : - (sheafForget.{_, _, _, _, max u₁ v₁} (D := D) J).IsRightAdjoint := - (adjunctionToTypes J (Adjunction.ofIsRightAdjoint (forget D))).isRightAdjoint +example [(forget D).IsRightAdjoint] : + (sheafForget.{_, _, _, _, max u₁ v₁} (D := D) J).IsRightAdjoint := by infer_instance end ForgetToType diff --git a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean index bd98aeb91cd38..e7d77d4762d54 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/LocallySurjective.lean @@ -53,7 +53,7 @@ lemma regularTopology.isLocallySurjective_iff [Preregular C] {F G : Cᵒᵖ ⥤ rw [regularTopology.mem_sieves_iff_hasEffectiveEpi] exact ⟨X', π, h, h'⟩ -lemma extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes [FinitaryPreExtensive C] +lemma extensiveTopology.surjective_of_isLocallySurjective_sheaf_of_types [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] (h : Presheaf.IsLocallySurjective (extensiveTopology C) f) {X : C} : Function.Surjective (f.app (op X)) := by @@ -84,13 +84,17 @@ lemma extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes [Finitary erw [IsLimit.map_π] rfl +@[deprecated (since := "2024-11-26")] +alias extensiveTopology.surjective_of_isLocallySurjective_sheafOfTypes := + extensiveTopology.surjective_of_isLocallySurjective_sheaf_of_types + lemma extensiveTopology.presheafIsLocallySurjective_iff [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ D} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] [PreservesFiniteProducts (forget D)] : Presheaf.IsLocallySurjective (extensiveTopology C) f ↔ ∀ (X : C), Function.Surjective (f.app (op X)) := by constructor · rw [Presheaf.isLocallySurjective_iff_whisker_forget (J := extensiveTopology C)] - exact fun h _ ↦ surjective_of_isLocallySurjective_sheafOfTypes (whiskerRight f (forget D)) h + exact fun h _ ↦ surjective_of_isLocallySurjective_sheaf_of_types (whiskerRight f (forget D)) h · intro h refine ⟨fun {X} y ↦ ?_⟩ obtain ⟨x, hx⟩ := h X y @@ -104,7 +108,7 @@ lemma extensiveTopology.isLocallySurjective_iff [FinitaryExtensive C] ∀ (X : C), Function.Surjective (f.val.app (op X)) := extensiveTopology.presheafIsLocallySurjective_iff _ f.val -lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryPreExtensive C] +lemma regularTopology.isLocallySurjective_sheaf_of_types [Preregular C] [FinitaryPreExtensive C] {F G : Cᵒᵖ ⥤ Type w} (f : F ⟶ G) [PreservesFiniteProducts F] [PreservesFiniteProducts G] (h : Presheaf.IsLocallySurjective (coherentTopology C) f) : Presheaf.IsLocallySurjective (regularTopology C) f where @@ -136,6 +140,9 @@ lemma regularTopology.isLocallySurjective_sheafOfTypes [Preregular C] [FinitaryP · change G.map _ (G.map _ _) = _ simp only [← FunctorToTypes.map_comp_apply, ← op_comp, Sigma.ι_desc] +@[deprecated (since := "2024-11-26")] alias regularTopology.isLocallySurjective_sheafOfTypes := +regularTopology.isLocallySurjective_sheaf_of_types + lemma coherentTopology.presheafIsLocallySurjective_iff {F G : Cᵒᵖ ⥤ D} (f : F ⟶ G) [Preregular C] [FinitaryPreExtensive C] [PreservesFiniteProducts F] [PreservesFiniteProducts G] [PreservesFiniteProducts (forget D)] : @@ -144,7 +151,7 @@ lemma coherentTopology.presheafIsLocallySurjective_iff {F G : Cᵒᵖ ⥤ D} (f constructor · rw [Presheaf.isLocallySurjective_iff_whisker_forget, Presheaf.isLocallySurjective_iff_whisker_forget (J := regularTopology C)] - exact regularTopology.isLocallySurjective_sheafOfTypes _ + exact regularTopology.isLocallySurjective_sheaf_of_types _ · refine Presheaf.isLocallySurjective_of_le (J := regularTopology C) ?_ _ rw [← extensive_regular_generate_coherent] exact (Coverage.gi _).gc.monotone_l le_sup_right diff --git a/Mathlib/CategoryTheory/Sites/Continuous.lean b/Mathlib/CategoryTheory/Sites/Continuous.lean index 91469d869e53b..ef0c9ef232567 100644 --- a/Mathlib/CategoryTheory/Sites/Continuous.lean +++ b/Mathlib/CategoryTheory/Sites/Continuous.lean @@ -115,31 +115,35 @@ abbrev PreservesOneHypercovers := /-- A functor `F` is continuous if the precomposition with `F.op` sends sheaves of `Type t` to sheaves. -/ class IsContinuous : Prop where - op_comp_isSheafOfTypes (G : SheafOfTypes.{t} K) : Presieve.IsSheaf J (F.op ⋙ G.val) + op_comp_isSheaf_of_types (G : Sheaf K (Type t)) : Presieve.IsSheaf J (F.op ⋙ G.val) -lemma op_comp_isSheafOfTypes [Functor.IsContinuous.{t} F J K] (G : SheafOfTypes.{t} K) : +lemma op_comp_isSheaf_of_types [Functor.IsContinuous.{t} F J K] (G : Sheaf K (Type t)) : Presieve.IsSheaf J (F.op ⋙ G.val) := - Functor.IsContinuous.op_comp_isSheafOfTypes _ + Functor.IsContinuous.op_comp_isSheaf_of_types _ + +@[deprecated (since := "2024-11-26")] alias op_comp_isSheafOfTypes := op_comp_isSheaf_of_types lemma op_comp_isSheaf [Functor.IsContinuous.{t} F J K] (G : Sheaf K A) : Presheaf.IsSheaf J (F.op ⋙ G.val) := - fun T => F.op_comp_isSheafOfTypes J K ⟨_, G.cond T⟩ + fun T => F.op_comp_isSheaf_of_types J K ⟨_, (isSheaf_iff_isSheaf_of_type _ _).2 (G.cond T)⟩ lemma isContinuous_of_iso {F₁ F₂ : C ⥤ D} (e : F₁ ≅ F₂) (J : GrothendieckTopology C) (K : GrothendieckTopology D) [Functor.IsContinuous.{t} F₁ J K] : Functor.IsContinuous.{t} F₂ J K where - op_comp_isSheafOfTypes G := + op_comp_isSheaf_of_types G := Presieve.isSheaf_iso J (isoWhiskerRight (NatIso.op e.symm) _) - (F₁.op_comp_isSheafOfTypes J K G) + (F₁.op_comp_isSheaf_of_types J K G) instance isContinuous_id : Functor.IsContinuous.{w} (𝟭 C) J J where - op_comp_isSheafOfTypes G := G.2 + op_comp_isSheaf_of_types G := (isSheaf_iff_isSheaf_of_type _ _).1 G.2 lemma isContinuous_comp (F₁ : C ⥤ D) (F₂ : D ⥤ E) (J : GrothendieckTopology C) (K : GrothendieckTopology D) (L : GrothendieckTopology E) [Functor.IsContinuous.{t} F₁ J K] [Functor.IsContinuous.{t} F₂ K L] : Functor.IsContinuous.{t} (F₁ ⋙ F₂) J L where - op_comp_isSheafOfTypes G := F₁.op_comp_isSheafOfTypes J K ⟨_, F₂.op_comp_isSheafOfTypes K L G⟩ + op_comp_isSheaf_of_types G := + F₁.op_comp_isSheaf_of_types J K + ⟨_,(isSheaf_iff_isSheaf_of_type _ _).2 (F₂.op_comp_isSheaf_of_types K L G)⟩ lemma isContinuous_comp' {F₁ : C ⥤ D} {F₂ : D ⥤ E} {F₁₂ : C ⥤ E} (e : F₁ ⋙ F₂ ≅ F₁₂) (J : GrothendieckTopology C) @@ -163,9 +167,9 @@ lemma op_comp_isSheaf_of_preservesOneHypercovers lemma isContinuous_of_preservesOneHypercovers [PreservesOneHypercovers.{w} F J K] [GrothendieckTopology.IsGeneratedByOneHypercovers.{w} J] : IsContinuous.{t} F J K where - op_comp_isSheafOfTypes := by + op_comp_isSheaf_of_types := by rintro ⟨P, hP⟩ - rw [← isSheaf_iff_isSheaf_of_type] at hP ⊢ + rw [← isSheaf_iff_isSheaf_of_type] exact F.op_comp_isSheaf_of_preservesOneHypercovers J K P hP end diff --git a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean index 54301e987ef0b..67fded0c9b4c2 100644 --- a/Mathlib/CategoryTheory/Sites/CoverPreserving.lean +++ b/Mathlib/CategoryTheory/Sites/CoverPreserving.lean @@ -75,13 +75,13 @@ This is actually stronger than merely preserving compatible families because of -- Porting note (https://github.com/leanprover-community/mathlib4/issues/5171): linter not ported yet @[nolint has_nonempty_instance] structure CompatiblePreserving (K : GrothendieckTopology D) (G : C ⥤ D) : Prop where compatible : - ∀ (ℱ : SheafOfTypes.{w} K) {Z} {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} + ∀ (ℱ : Sheaf K (Type w)) {Z} {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} (_ : x.Compatible) {Y₁ Y₂} {X} (f₁ : X ⟶ G.obj Y₁) (f₂ : X ⟶ G.obj Y₂) {g₁ : Y₁ ⟶ Z} {g₂ : Y₂ ⟶ Z} (hg₁ : T g₁) (hg₂ : T g₂) (_ : f₁ ≫ G.map g₁ = f₂ ≫ G.map g₂), ℱ.val.map f₁.op (x g₁ hg₁) = ℱ.val.map f₂.op (x g₂ hg₂) section -variable {J K} {G : C ⥤ D} (hG : CompatiblePreserving.{w} K G) (ℱ : SheafOfTypes.{w} K) {Z : C} +variable {J K} {G : C ⥤ D} (hG : CompatiblePreserving.{w} K G) (ℱ : Sheaf K (Type w)) {Z : C} variable {T : Presieve Z} {x : FamilyOfElements (G.op ⋙ ℱ.val) T} (h : x.Compatible) include hG h @@ -163,14 +163,15 @@ This result is basically . -/ lemma Functor.isContinuous_of_coverPreserving (hF₁ : CompatiblePreserving.{w} K F) (hF₂ : CoverPreserving J K F) : Functor.IsContinuous.{w} F J K where - op_comp_isSheafOfTypes G X S hS x hx := by + op_comp_isSheaf_of_types G X S hS x hx := by apply exists_unique_of_exists_of_unique - · have H := G.2 _ (hF₂.cover_preserve hS) + · have H := (isSheaf_iff_isSheaf_of_type _ _).1 G.2 _ (hF₂.cover_preserve hS) exact ⟨H.amalgamate (x.functorPushforward F) (hx.functorPushforward hF₁), fun V f hf => (H.isAmalgamation (hx.functorPushforward hF₁) (F.map f) _).trans (hF₁.apply_map _ hx hf)⟩ · intro y₁ y₂ hy₁ hy₂ - apply (Presieve.isSeparated_of_isSheaf _ _ G.cond _ (hF₂.cover_preserve hS)).ext + apply (Presieve.isSeparated_of_isSheaf _ _ ((isSheaf_iff_isSheaf_of_type _ _).1 G.2) _ + (hF₂.cover_preserve hS)).ext rintro Y _ ⟨Z, g, h, hg, rfl⟩ dsimp simp only [Functor.map_comp, types_comp_apply] diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean index fb9eb889b3ee1..5e24ab6fe6ff6 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite.lean @@ -115,9 +115,10 @@ variable {K} variable {A : Type*} [Category A] (G : C ⥤ D) -- this is not marked with `@[ext]` because `H` can not be inferred from the type -theorem ext [G.IsCoverDense K] (ℱ : SheafOfTypes K) (X : D) {s t : ℱ.val.obj (op X)} +theorem ext [G.IsCoverDense K] (ℱ : Sheaf K (Type _)) (X : D) {s t : ℱ.val.obj (op X)} (h : ∀ ⦃Y : C⦄ (f : G.obj Y ⟶ X), ℱ.val.map f.op s = ℱ.val.map f.op t) : s = t := by - apply (ℱ.cond (Sieve.coverByImage G X) (G.is_cover_of_isCoverDense K X)).isSeparatedFor.ext + apply ((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ.cond + (Sieve.coverByImage G X) (G.is_cover_of_isCoverDense K X)).isSeparatedFor.ext rintro Y _ ⟨Z, f₁, f₂, ⟨rfl⟩⟩ simp [h f₂] @@ -156,7 +157,7 @@ theorem sheaf_eq_amalgamation (ℱ : Sheaf K A) {X : A} {U : D} {T : Sieve U} (h namespace Types -variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : SheafOfTypes.{v} K} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) +variable {ℱ : Dᵒᵖ ⥤ Type v} {ℱ' : Sheaf K (Type v)} (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) theorem naturality_apply [G.IsLocallyFull K] {X Y : C} (i : G.obj X ⟶ G.obj Y) (x) : ℱ'.1.map i.op (α.app _ x) = α.app _ (ℱ.map i.op x) := by @@ -212,13 +213,14 @@ theorem pushforwardFamily_compatible {X} (x : ℱ.obj (op X)) : /-- (Implementation). The morphism `ℱ(X) ⟶ ℱ'(X)` given by gluing the `pushforwardFamily`. -/ noncomputable def appHom (X : D) : ℱ.obj (op X) ⟶ ℱ'.val.obj (op X) := fun x => - (ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).amalgamate (pushforwardFamily α x) - (pushforwardFamily_compatible α x) + ((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ'.cond _ + (G.is_cover_of_isCoverDense _ X)).amalgamate (pushforwardFamily α x) + (pushforwardFamily_compatible α x) @[simp] theorem appHom_restrict {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) (x) : ℱ'.val.map f (appHom α X x) = α.app (op Y) (ℱ.map f x) := - ((ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).valid_glue + (((isSheaf_iff_isSheaf_of_type _ _ ).1 ℱ'.cond _ (G.is_cover_of_isCoverDense _ X)).valid_glue (pushforwardFamily_compatible α x) f.unop (Presieve.in_coverByImage G f.unop)).trans (pushforwardFamily_apply _ _ _) @@ -232,7 +234,7 @@ theorem appHom_valid_glue {X : D} {Y : C} (f : op X ⟶ op (G.obj Y)) : (Implementation). The maps given in `appIso` is inverse to each other and gives a `ℱ(X) ≅ ℱ'(X)`. -/ @[simps] -noncomputable def appIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) +noncomputable def appIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) (X : D) : ℱ.val.obj (op X) ≅ ℱ'.val.obj (op X) where hom := appHom i.hom X inv := appHom i.inv X @@ -267,7 +269,7 @@ where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between presheaves. -/ @[simps!] -noncomputable def presheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : +noncomputable def presheafIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : ℱ.val ≅ ℱ'.val := NatIso.ofComponents (fun X => appIso i (unop X)) @(presheafHom i.hom).naturality @@ -277,7 +279,7 @@ where `G` is locally-full and cover-dense, and `ℱ, ℱ'` are sheaves, we may obtain a natural isomorphism between sheaves. -/ @[simps] -noncomputable def sheafIso {ℱ ℱ' : SheafOfTypes.{v} K} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : +noncomputable def sheafIso {ℱ ℱ' : Sheaf K (Type v)} (i : G.op ⋙ ℱ.val ≅ G.op ⋙ ℱ'.val) : ℱ ≅ ℱ' where hom := ⟨(presheafIso i).hom⟩ inv := ⟨(presheafIso i).inv⟩ @@ -398,7 +400,9 @@ theorem sheafHom_restrict_eq (α : G.op ⋙ ℱ ⟶ G.op ⋙ ℱ'.val) : simp only [Category.assoc] congr 1 have := naturality_apply (G := G) (ℱ := ℱ ⋙ coyoneda.obj (op <| (G.op ⋙ ℱ).obj X)) - (ℱ' := ⟨_, ℱ'.2 ((G.op ⋙ ℱ).obj X)⟩) (whiskerRight α (coyoneda.obj _)) hf.some.map (𝟙 _) + (ℱ' := ⟨_, Presheaf.isSheaf_comp_of_isSheaf K ℱ'.val + (coyoneda.obj (op ((G.op ⋙ ℱ).obj X))) ℱ'.cond⟩) + (whiskerRight α (coyoneda.obj _)) hf.some.map (𝟙 _) simpa using this variable (G) diff --git a/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean index 96bf30d33b13f..2988aa4190f18 100644 --- a/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean +++ b/Mathlib/CategoryTheory/Sites/LocallyFullyFaithful.lean @@ -103,20 +103,22 @@ variable {K} variable {A : Type*} [Category A] (G : C ⥤ D) theorem IsLocallyFull.ext [G.IsLocallyFull K] - (ℱ : SheafOfTypes K) {X Y : C} (i : G.obj X ⟶ G.obj Y) + (ℱ : Sheaf K (Type _)) {X Y : C} (i : G.obj X ⟶ G.obj Y) {s t : ℱ.val.obj (op (G.obj X))} (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X) (f : Z ⟶ Y), G.map f = G.map j ≫ i → ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by - apply (ℱ.cond _ (G.functorPushforward_imageSieve_mem K i)).isSeparatedFor.ext + apply (((isSheaf_iff_isSheaf_of_type _ _).1 ℱ.cond) _ + (G.functorPushforward_imageSieve_mem K i)).isSeparatedFor.ext rintro Z _ ⟨W, iWX, iZW, ⟨iWY, e⟩, rfl⟩ simp [h iWX iWY e] -theorem IsLocallyFaithful.ext [G.IsLocallyFaithful K] (ℱ : SheafOfTypes K) +theorem IsLocallyFaithful.ext [G.IsLocallyFaithful K] (ℱ : Sheaf K (Type _)) {X Y : C} (i₁ i₂ : X ⟶ Y) (e : G.map i₁ = G.map i₂) {s t : ℱ.val.obj (op (G.obj X))} (h : ∀ ⦃Z : C⦄ (j : Z ⟶ X), j ≫ i₁ = j ≫ i₂ → ℱ.1.map (G.map j).op s = ℱ.1.map (G.map j).op t) : s = t := by - apply (ℱ.cond _ (G.functorPushforward_equalizer_mem K i₁ i₂ e)).isSeparatedFor.ext + apply (((isSheaf_iff_isSheaf_of_type _ _).1 ℱ.cond) _ + (G.functorPushforward_equalizer_mem K i₁ i₂ e)).isSeparatedFor.ext rintro Z _ ⟨W, iWX, iZW, hiWX, rfl⟩ simp [h iWX hiWX] diff --git a/Mathlib/CategoryTheory/Sites/Sheaf.lean b/Mathlib/CategoryTheory/Sites/Sheaf.lean index 9ef621b55f223..6900396ec52a3 100644 --- a/Mathlib/CategoryTheory/Sites/Sheaf.lean +++ b/Mathlib/CategoryTheory/Sites/Sheaf.lean @@ -376,12 +376,6 @@ theorem Sheaf.Hom.mono_of_presheaf_mono {F G : Sheaf J A} (f : F ⟶ G) [h : Mon instance Sheaf.Hom.epi_of_presheaf_epi {F G : Sheaf J A} (f : F ⟶ G) [h : Epi f.1] : Epi f := (sheafToPresheaf J A).epi_of_epi_map h -/-- The sheaf of sections guaranteed by the sheaf condition. -/ -@[simps] -def sheafOver {A : Type u₂} [Category.{v₂} A] {J : GrothendieckTopology C} (ℱ : Sheaf J A) (E : A) : - SheafOfTypes J := - ⟨ℱ.val ⋙ coyoneda.obj (op E), ℱ.cond E⟩ - theorem isSheaf_iff_isSheaf_of_type (P : Cᵒᵖ ⥤ Type w) : Presheaf.IsSheaf J P ↔ Presieve.IsSheaf J P := by constructor @@ -402,27 +396,39 @@ theorem isSheaf_iff_isSheaf_of_type (P : Cᵒᵖ ⥤ Type w) : rw [Presieve.IsSheafFor.valid_glue _ _ _ hf, ← hy _ hf] rfl +/-- The sheaf of sections guaranteed by the sheaf condition. -/ +@[simps] +def sheafOver {A : Type u₂} [Category.{v₂} A] {J : GrothendieckTopology C} (ℱ : Sheaf J A) (E : A) : + Sheaf J (Type _) where + val := ℱ.val ⋙ coyoneda.obj (op E) + cond := by + rw [isSheaf_iff_isSheaf_of_type] + exact ℱ.cond E + variable {J} in lemma Presheaf.IsSheaf.isSheafFor {P : Cᵒᵖ ⥤ Type w} (hP : Presheaf.IsSheaf J P) {X : C} (S : Sieve X) (hS : S ∈ J X) : Presieve.IsSheafFor P S.arrows := by rw [isSheaf_iff_isSheaf_of_type] at hP exact hP S hS -/-- The category of sheaves taking values in Type is the same as the category of set-valued sheaves. +variable {A} in +lemma Presheaf.isSheaf_bot (P : Cᵒᵖ ⥤ A) : IsSheaf ⊥ P := fun _ ↦ Presieve.isSheaf_bot + +/-- +The category of sheaves on the bottom (trivial) Grothendieck topology is +equivalent to the category of presheaves. -/ @[simps] -def sheafEquivSheafOfTypes : Sheaf J (Type w) ≌ SheafOfTypes J where - functor := - { obj := fun S => ⟨S.val, (isSheaf_iff_isSheaf_of_type _ _).1 S.2⟩ - map := fun f => ⟨f.val⟩ } +def sheafBotEquivalence : Sheaf (⊥ : GrothendieckTopology C) A ≌ Cᵒᵖ ⥤ A where + functor := sheafToPresheaf _ _ inverse := - { obj := fun S => ⟨S.val, (isSheaf_iff_isSheaf_of_type _ _).2 S.2⟩ - map := fun f => ⟨f.val⟩ } - unitIso := NatIso.ofComponents fun _ => Iso.refl _ - counitIso := NatIso.ofComponents fun _ => Iso.refl _ + { obj := fun P => ⟨P, Presheaf.isSheaf_bot P⟩ + map := fun f => ⟨f⟩ } + unitIso := Iso.refl _ + counitIso := Iso.refl _ instance : Inhabited (Sheaf (⊥ : GrothendieckTopology C) (Type w)) := - ⟨(sheafEquivSheafOfTypes _).inverse.obj default⟩ + ⟨(sheafBotEquivalence _).inverse.obj ((Functor.const _).obj default)⟩ variable {J} {A} diff --git a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean index ec2ba32663333..9819869ea4a96 100644 --- a/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean +++ b/Mathlib/CategoryTheory/Sites/SheafOfTypes.lean @@ -198,74 +198,4 @@ theorem forallYonedaIsSheaf_iff_colimit (S : Sieve X) : end Sieve -variable {C : Type u} [Category.{v} C] -variable (J : GrothendieckTopology C) - -/-- The category of sheaves on a grothendieck topology. -/ -structure SheafOfTypes (J : GrothendieckTopology C) : Type max u v (w + 1) where - /-- the underlying presheaf -/ - val : Cᵒᵖ ⥤ Type w - /-- the condition that the presheaf is a sheaf -/ - cond : Presieve.IsSheaf J val - -namespace SheafOfTypes - -variable {J} - -/-- Morphisms between sheaves of types are just morphisms between the underlying presheaves. -/ -@[ext] -structure Hom (X Y : SheafOfTypes J) where - /-- a morphism between the underlying presheaves -/ - val : X.val ⟶ Y.val - -@[simps] -instance : Category (SheafOfTypes J) where - Hom := Hom - id _ := ⟨𝟙 _⟩ - comp f g := ⟨f.val ≫ g.val⟩ - id_comp _ := Hom.ext <| id_comp _ - comp_id _ := Hom.ext <| comp_id _ - assoc _ _ _ := Hom.ext <| assoc _ _ _ - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11041): we need to restate the `ext` lemma in terms of the categorical morphism. --- not just the underlying structure. --- It would be nice if this boilerplate weren't necessary. -@[ext] -theorem Hom.ext' {X Y : SheafOfTypes J} (f g : X ⟶ Y) (w : f.val = g.val) : f = g := - Hom.ext w - --- Let's make the inhabited linter happy... -instance (X : SheafOfTypes J) : Inhabited (Hom X X) := - ⟨𝟙 X⟩ - -end SheafOfTypes - -/-- The inclusion functor from sheaves to presheaves. -/ -@[simps] -def sheafOfTypesToPresheaf : SheafOfTypes J ⥤ Cᵒᵖ ⥤ Type w where - obj := SheafOfTypes.val - map f := f.val - map_id _ := rfl - map_comp _ _ := rfl - -instance : (sheafOfTypesToPresheaf J).Full where map_surjective f := ⟨⟨f⟩, rfl⟩ - -instance : (sheafOfTypesToPresheaf J).Faithful where - -/-- -The category of sheaves on the bottom (trivial) grothendieck topology is equivalent to the category -of presheaves. --/ -@[simps] -def sheafOfTypesBotEquiv : SheafOfTypes (⊥ : GrothendieckTopology C) ≌ Cᵒᵖ ⥤ Type w where - functor := sheafOfTypesToPresheaf _ - inverse := - { obj := fun P => ⟨P, Presieve.isSheaf_bot⟩ - map := fun f => ⟨f⟩ } - unitIso := Iso.refl _ - counitIso := Iso.refl _ - -instance : Inhabited (SheafOfTypes (⊥ : GrothendieckTopology C)) := - ⟨sheafOfTypesBotEquiv.inverse.obj ((Functor.const _).obj PUnit)⟩ - end CategoryTheory diff --git a/Mathlib/CategoryTheory/Sites/Types.lean b/Mathlib/CategoryTheory/Sites/Types.lean index cfa559d44c11a..65430a4c42150 100644 --- a/Mathlib/CategoryTheory/Sites/Types.lean +++ b/Mathlib/CategoryTheory/Sites/Types.lean @@ -45,23 +45,31 @@ theorem generate_discretePresieve_mem (α : Type u) : Sieve.generate (discretePresieve α) ∈ typesGrothendieckTopology α := fun x => ⟨PUnit, id, fun _ => x, ⟨PUnit.unit, fun _ => Subsingleton.elim _ _⟩, rfl⟩ -open Presieve - -theorem isSheaf_yoneda' {α : Type u} : IsSheaf typesGrothendieckTopology (yoneda.obj α) := +/-- The sheaf condition for `yoneda'`. -/ +theorem Presieve.isSheaf_yoneda' {α : Type u} : + Presieve.IsSheaf typesGrothendieckTopology (yoneda.obj α) := fun β _ hs x hx => ⟨fun y => x _ (hs y) PUnit.unit, fun γ f h => funext fun z => by convert congr_fun (hx (𝟙 _) (fun _ => z) (hs <| f z) h rfl) PUnit.unit using 1, fun f hf => funext fun y => by convert congr_fun (hf _ (hs y)) PUnit.unit⟩ +/-- The sheaf condition for `yoneda'`. -/ +theorem Presheaf.isSheaf_yoneda' {α : Type u} : + Presheaf.IsSheaf typesGrothendieckTopology (yoneda.obj α) := by + rw [isSheaf_iff_isSheaf_of_type] + exact Presieve.isSheaf_yoneda' + +@[deprecated (since := "2024-11-26")] alias isSheaf_yoneda' := Presieve.isSheaf_yoneda' + /-- The yoneda functor that sends a type to a sheaf over the category of types. -/ @[simps] -def yoneda' : Type u ⥤ SheafOfTypes typesGrothendieckTopology where - obj α := ⟨yoneda.obj α, isSheaf_yoneda'⟩ +def yoneda' : Type u ⥤ Sheaf typesGrothendieckTopology (Type u) where + obj α := ⟨yoneda.obj α, Presheaf.isSheaf_yoneda'⟩ map f := ⟨yoneda.map f⟩ @[simp] -theorem yoneda'_comp : yoneda'.{u} ⋙ sheafOfTypesToPresheaf _ = yoneda := +theorem yoneda'_comp : yoneda'.{u} ⋙ sheafToPresheaf _ _ = yoneda := rfl open Opposite @@ -71,6 +79,8 @@ a map `P(α) → (α → P(*))` for all type `α`. -/ def eval (P : Type uᵒᵖ ⥤ Type u) (α : Type u) (s : P.obj (op α)) (x : α) : P.obj (op PUnit) := P.map (↾fun _ => x).op s +open Presieve + /-- Given a sheaf `S` on the category of types, construct a map `(α → S(*)) → S(α)` that is inverse to `eval`. -/ noncomputable def typesGlue (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) @@ -103,10 +113,11 @@ theorem typesGlue_eval {S hs α} (s) : typesGlue.{u} S hs α (eval S α s) = s : /-- Given a sheaf `S`, construct an equivalence `S(α) ≃ (α → S(*))`. -/ @[simps] -noncomputable def evalEquiv (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) +noncomputable def evalEquiv (S : Type uᵒᵖ ⥤ Type u) + (hs : Presheaf.IsSheaf typesGrothendieckTopology S) (α : Type u) : S.obj (op α) ≃ (α → S.obj (op PUnit)) where toFun := eval S α - invFun := typesGlue S hs α + invFun := typesGlue S ((isSheaf_iff_isSheaf_of_type _ _ ).1 hs) α left_inv := typesGlue_eval right_inv := eval_typesGlue @@ -116,31 +127,32 @@ theorem eval_map (S : Type uᵒᵖ ⥤ Type u) (α β) (f : β ⟶ α) (s x) : /-- Given a sheaf `S`, construct an isomorphism `S ≅ [-, S(*)]`. -/ @[simps!] -noncomputable def equivYoneda (S : Type uᵒᵖ ⥤ Type u) (hs : IsSheaf typesGrothendieckTopology S) : +noncomputable def equivYoneda (S : Type uᵒᵖ ⥤ Type u) + (hs : Presheaf.IsSheaf typesGrothendieckTopology S) : S ≅ yoneda.obj (S.obj (op PUnit)) := NatIso.ofComponents (fun α => Equiv.toIso <| evalEquiv S hs <| unop α) fun {α β} f => funext fun _ => funext fun _ => eval_map S (unop α) (unop β) f.unop _ _ /-- Given a sheaf `S`, construct an isomorphism `S ≅ [-, S(*)]`. -/ @[simps] -noncomputable def equivYoneda' (S : SheafOfTypes typesGrothendieckTopology) : +noncomputable def equivYoneda' (S : Sheaf typesGrothendieckTopology (Type u)) : S ≅ yoneda'.obj (S.1.obj (op PUnit)) where hom := ⟨(equivYoneda S.1 S.2).hom⟩ inv := ⟨(equivYoneda S.1 S.2).inv⟩ hom_inv_id := by ext1; apply (equivYoneda S.1 S.2).hom_inv_id inv_hom_id := by ext1; apply (equivYoneda S.1 S.2).inv_hom_id -theorem eval_app (S₁ S₂ : SheafOfTypes.{u} typesGrothendieckTopology) (f : S₁ ⟶ S₂) (α : Type u) +theorem eval_app (S₁ S₂ : Sheaf typesGrothendieckTopology (Type u)) (f : S₁ ⟶ S₂) (α : Type u) (s : S₁.1.obj (op α)) (x : α) : eval S₂.1 α (f.val.app (op α) s) x = f.val.app (op PUnit) (eval S₁.1 α s x) := (congr_fun (f.val.naturality (↾fun _ : PUnit => x).op) s).symm /-- `yoneda'` induces an equivalence of category between `Type u` and -`SheafOfTypes typesGrothendieckTopology`. -/ +`Sheaf typesGrothendieckTopology (Type u)`. -/ @[simps!] -noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology where +noncomputable def typeEquiv : Type u ≌ Sheaf typesGrothendieckTopology (Type u) where functor := yoneda' - inverse := sheafOfTypesToPresheaf _ ⋙ (evaluation _ _).obj (op PUnit) + inverse := sheafToPresheaf _ _ ⋙ (evaluation _ _).obj (op PUnit) unitIso := NatIso.ofComponents (fun _α => -- α ≅ PUnit ⟶ α { hom := fun x _ => x @@ -149,9 +161,11 @@ noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology inv_hom_id := funext fun _ => funext fun y => PUnit.casesOn y rfl }) fun _ => rfl counitIso := Iso.symm <| - NatIso.ofComponents (fun S => equivYoneda' S) fun {S₁ S₂} f => - SheafOfTypes.Hom.ext <| NatTrans.ext <| - funext fun α => funext fun s => funext fun x => eval_app S₁ S₂ f (unop α) s x + NatIso.ofComponents (fun S => equivYoneda' S) (fun {S₁ S₂} f => by + ext ⟨α⟩ s + dsimp at s ⊢ + ext x + exact eval_app S₁ S₂ f α s x) functor_unitIso_comp X := by ext1 apply yonedaEquiv.injective @@ -159,7 +173,7 @@ noncomputable def typeEquiv : Type u ≌ SheafOfTypes typesGrothendieckTopology erw [typesGlue_eval] instance subcanonical_typesGrothendieckTopology : typesGrothendieckTopology.{u}.Subcanonical := - GrothendieckTopology.Subcanonical.of_isSheaf_yoneda_obj _ fun _ => isSheaf_yoneda' + GrothendieckTopology.Subcanonical.of_isSheaf_yoneda_obj _ fun _ => Presieve.isSheaf_yoneda' theorem typesGrothendieckTopology_eq_canonical : typesGrothendieckTopology.{u} = Sheaf.canonicalTopology (Type u) := by @@ -167,7 +181,7 @@ theorem typesGrothendieckTopology_eq_canonical : refine ⟨yoneda.obj (ULift Bool), ⟨_, rfl⟩, GrothendieckTopology.ext ?_⟩ funext α ext S - refine ⟨fun hs x => ?_, fun hs β f => isSheaf_yoneda' _ fun y => hs _⟩ + refine ⟨fun hs x => ?_, fun hs β f => Presieve.isSheaf_yoneda' _ fun y => hs _⟩ by_contra hsx have : (fun _ => ULift.up true) = fun _ => ULift.up false := (hs PUnit fun _ => x).isSeparatedFor.ext From ec29ecebbee207d4d121b387652d5b846384dc56 Mon Sep 17 00:00:00 2001 From: Andrew Yang Date: Tue, 26 Nov 2024 23:28:21 +0000 Subject: [PATCH 205/250] feat(Topology/Algebra): add `IsModuleTopology.continuous_of_ringHom` (#19515) --- Mathlib/Topology/Algebra/Module/ModuleTopology.lean | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean index 4de1767692be2..e5366ba9ea100 100644 --- a/Mathlib/Topology/Algebra/Module/ModuleTopology.lean +++ b/Mathlib/Topology/Algebra/Module/ModuleTopology.lean @@ -292,6 +292,16 @@ theorem continuous_neg (C : Type*) [AddCommGroup C] [Module R C] [TopologicalSpa haveI : ContinuousAdd C := IsModuleTopology.toContinuousAdd R C continuous_of_linearMap (LinearEquiv.neg R).toLinearMap +@[fun_prop, continuity] +theorem continuous_of_ringHom {R A B} [CommSemiring R] [Semiring A] [Algebra R A] [Semiring B] + [TopologicalSpace R] [TopologicalSpace A] [IsModuleTopology R A] [TopologicalSpace B] + [TopologicalSemiring B] + (φ : A →+* B) (hφ : Continuous (φ.comp (algebraMap R A))) : Continuous φ := by + let inst := Module.compHom B (φ.comp (algebraMap R A)) + let φ' : A →ₗ[R] B := ⟨φ, fun r m ↦ by simp [Algebra.smul_def]; rfl⟩ + have : ContinuousSMul R B := ⟨(hφ.comp continuous_fst).mul continuous_snd⟩ + exact continuous_of_linearMap φ' + end function end IsModuleTopology From 8c20b1c7204155a2b5a58338267c56d0457a3556 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 11:15:51 +1100 Subject: [PATCH 206/250] increase heartbeats --- MathlibTest/linear_combination.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/MathlibTest/linear_combination.lean b/MathlibTest/linear_combination.lean index b947d8190146e..96eecab1651c1 100644 --- a/MathlibTest/linear_combination.lean +++ b/MathlibTest/linear_combination.lean @@ -521,7 +521,7 @@ typeclass inference is demanded by the lemmas it orchestrates. This example too (and 73 ms on a good laptop) on an implementation with "minimal" typeclasses everywhere, e.g. lots of `CovariantClass`/`ContravariantClass`, and takes 206 heartbeats (10 ms on a good laptop) on the implementation at the time of joining Mathlib (November 2024). -/ -set_option maxHeartbeats 1100 in +set_option maxHeartbeats 1200 in example {a b : ℝ} (h : a < b) : 0 < b - a := by linear_combination (norm := skip) h exact test_sorry From 69959750df8adb05e7a31e9aa0a8fbd850188956 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 27 Nov 2024 02:02:50 +0000 Subject: [PATCH 207/250] feat: add `NonUnital{Star}Algebra.elemental` subalgebras (#18615) Define the `NonUnital{Star}Algebra.elemental` as `(adjoin R {x}).topologicalClosure` to match the unital versions. - [ ] depends on: #18612 - [ ] depends on: #18613 --- .../Topology/Algebra/NonUnitalAlgebra.lean | 61 +++++++++++++++++ .../Algebra/NonUnitalStarAlgebra.lean | 68 +++++++++++++++++++ 2 files changed, 129 insertions(+) diff --git a/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean b/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean index 91c54d048c2f2..e450112e92720 100644 --- a/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean +++ b/Mathlib/Topology/Algebra/NonUnitalAlgebra.lean @@ -77,3 +77,64 @@ abbrev nonUnitalCommRingTopologicalClosure [T2Space A] (s : NonUnitalSubalgebra end Ring end NonUnitalSubalgebra + +namespace NonUnitalAlgebra + +open NonUnitalSubalgebra + +variable (R : Type*) {A : Type*} [CommSemiring R] [NonUnitalSemiring A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] +variable [TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A] + +/-- The topological closure of the non-unital subalgebra generated by a single element. -/ +def elemental (x : A) : NonUnitalSubalgebra R A := + adjoin R {x} |>.topologicalClosure + +namespace elemental + +@[aesop safe apply (rule_sets := [SetLike])] +theorem self_mem (x : A) : x ∈ elemental R x := + le_topologicalClosure _ <| self_mem_adjoin_singleton R x + +variable {R} in +theorem le_of_mem {x : A} {s : NonUnitalSubalgebra R A} (hs : IsClosed (s : Set A)) (hx : x ∈ s) : + elemental R x ≤ s := + topologicalClosure_minimal _ (adjoin_le <| by simpa using hx) hs + +variable {R} in +theorem le_iff_mem {x : A} {s : NonUnitalSubalgebra R A} (hs : IsClosed (s : Set A)) : + elemental R x ≤ s ↔ x ∈ s := + ⟨fun h ↦ h (self_mem R x), fun h ↦ le_of_mem hs h⟩ + +instance isClosed (x : A) : IsClosed (elemental R x : Set A) := + isClosed_topologicalClosure _ + +instance [T2Space A] {x : A} : NonUnitalCommSemiring (elemental R x) := + nonUnitalCommSemiringTopologicalClosure _ + letI : NonUnitalCommSemiring (adjoin R {x}) := + NonUnitalAlgebra.adjoinNonUnitalCommSemiringOfComm R fun y hy z hz => by + rw [Set.mem_singleton_iff] at hy hz + rw [hy, hz] + fun _ _ => mul_comm _ _ + +instance {R A : Type*} [CommRing R] [NonUnitalRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] + [TopologicalSpace A] [TopologicalRing A] [ContinuousConstSMul R A] + [T2Space A] {x : A} : NonUnitalCommRing (elemental R x) where + mul_comm := mul_comm + +instance {A : Type*} [UniformSpace A] [CompleteSpace A] [NonUnitalSemiring A] + [TopologicalSemiring A] [Module R A] [IsScalarTower R A A] + [SMulCommClass R A A] [ContinuousConstSMul R A] (x : A) : + CompleteSpace (elemental R x) := + isClosed_closure.completeSpace_coe + +/-- The coercion from an elemental algebra to the full algebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : Topology.IsClosedEmbedding ((↑) : elemental R x → A) where + eq_induced := rfl + injective := Subtype.coe_injective + isClosed_range := by simpa using isClosed R x + +end elemental + +end NonUnitalAlgebra diff --git a/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean b/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean index 92452ae7c1b4a..5c74a83675c40 100644 --- a/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean +++ b/Mathlib/Topology/Algebra/NonUnitalStarAlgebra.lean @@ -81,3 +81,71 @@ abbrev nonUnitalCommRingTopologicalClosure [T2Space A] (s : NonUnitalStarSubalge end Ring end NonUnitalStarSubalgebra + +namespace NonUnitalStarAlgebra + +open NonUnitalStarSubalgebra + +variable (R : Type*) {A : Type*} [CommSemiring R] [StarRing R] [NonUnitalSemiring A] [StarRing A] +variable [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] +variable [TopologicalSpace A] [TopologicalSemiring A] [ContinuousConstSMul R A] [ContinuousStar A] + +/-- The topological closure of the non-unital star subalgebra generated by a single element. -/ +def elemental (x : A) : NonUnitalStarSubalgebra R A := + adjoin R {x} |>.topologicalClosure + +namespace elemental + +@[aesop safe apply (rule_sets := [SetLike])] +theorem self_mem (x : A) : x ∈ elemental R x := + le_topologicalClosure _ <| self_mem_adjoin_singleton R x + +@[aesop safe apply (rule_sets := [SetLike])] +theorem star_self_mem (x : A) : star x ∈ elemental R x := + le_topologicalClosure _ <| star_self_mem_adjoin_singleton R x + +variable {R} in +theorem le_of_mem {x : A} {s : NonUnitalStarSubalgebra R A} (hs : IsClosed (s : Set A)) + (hx : x ∈ s) : elemental R x ≤ s := + topologicalClosure_minimal _ (adjoin_le <| by simpa using hx) hs + +variable {R} in +theorem le_iff_mem {x : A} {s : NonUnitalStarSubalgebra R A} (hs : IsClosed (s : Set A)) : + elemental R x ≤ s ↔ x ∈ s := + ⟨fun h ↦ h (self_mem R x), fun h ↦ le_of_mem hs h⟩ + +instance isClosed (x : A) : IsClosed (elemental R x : Set A) := + isClosed_topologicalClosure _ + +instance [T2Space A] {x : A} [IsStarNormal x] : NonUnitalCommSemiring (elemental R x) := + nonUnitalCommSemiringTopologicalClosure _ <| by + letI : NonUnitalCommSemiring (adjoin R {x}) := by + refine NonUnitalStarAlgebra.adjoinNonUnitalCommSemiringOfComm R ?_ ?_ + all_goals + intro y hy z hz + rw [Set.mem_singleton_iff] at hy hz + rw [hy, hz] + exact (star_comm_self' x).symm + exact fun _ _ => mul_comm _ _ + +instance {R A : Type*} [CommRing R] [StarRing R] [NonUnitalRing A] [StarRing A] + [Module R A] [IsScalarTower R A A] [SMulCommClass R A A] [StarModule R A] + [TopologicalSpace A] [TopologicalRing A] [ContinuousConstSMul R A] [ContinuousStar A] + [T2Space A] {x : A} [IsStarNormal x] : NonUnitalCommRing (elemental R x) where + mul_comm := mul_comm + +instance {A : Type*} [UniformSpace A] [CompleteSpace A] [NonUnitalSemiring A] [StarRing A] + [TopologicalSemiring A] [ContinuousStar A] [Module R A] [IsScalarTower R A A] + [SMulCommClass R A A] [StarModule R A] [ContinuousConstSMul R A] (x : A) : + CompleteSpace (elemental R x) := + isClosed_closure.completeSpace_coe + +/-- The coercion from an elemental algebra to the full algebra is a `IsClosedEmbedding`. -/ +theorem isClosedEmbedding_coe (x : A) : Topology.IsClosedEmbedding ((↑) : elemental R x → A) where + eq_induced := rfl + injective := Subtype.coe_injective + isClosed_range := by simpa using isClosed R x + +end elemental + +end NonUnitalStarAlgebra From bdc1ed1f38009380f142b3500d72050c62faaf81 Mon Sep 17 00:00:00 2001 From: Jireh Loreaux Date: Wed, 27 Nov 2024 03:15:31 +0000 Subject: [PATCH 208/250] =?UTF-8?q?feat:=20approximate=20units=20in=20C?= =?UTF-8?q?=E2=8B=86-algebras=20(#18506)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This constructs the canonical approximate unit in an arbitrary C⋆-algebra. This approximate unit is the filter corresponding to the basis of sections `{x | a ≤ x} ∩ closedBall 0 1`, where `0 ≤ a` and `‖a‖ < 1`. --- Mathlib.lean | 1 + .../CStarAlgebra/ApproximateUnit.lean | 221 +++++++++++++++++- .../SpecialFunctions/PosPart.lean | 55 +++++ .../ContinuousFunctionalCalculus/PosPart.lean | 26 +++ 4 files changed, 292 insertions(+), 11 deletions(-) create mode 100644 Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean diff --git a/Mathlib.lean b/Mathlib.lean index 1fe86e9134e3b..0195cb1242e40 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1107,6 +1107,7 @@ import Mathlib.Analysis.CStarAlgebra.Module.Constructions import Mathlib.Analysis.CStarAlgebra.Module.Defs import Mathlib.Analysis.CStarAlgebra.Module.Synonym import Mathlib.Analysis.CStarAlgebra.Multiplier +import Mathlib.Analysis.CStarAlgebra.SpecialFunctions.PosPart import Mathlib.Analysis.CStarAlgebra.Spectrum import Mathlib.Analysis.CStarAlgebra.Unitization import Mathlib.Analysis.CStarAlgebra.lpSpace diff --git a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean index 89f1b4ea09125..49852fd88a32f 100644 --- a/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean +++ b/Mathlib/Analysis/CStarAlgebra/ApproximateUnit.lean @@ -5,15 +5,21 @@ Authors: Jireh Loreaux -/ import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Order import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Isometric +import Mathlib.Analysis.CStarAlgebra.SpecialFunctions.PosPart import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.Rpow +import Mathlib.Topology.ApproximateUnit -/-! # Positive contractions in a C⋆-algebra form an approximate unit +/-! # Nonnegative contractions in a C⋆-algebra form an approximate unit -This file will show (although it does not yet) that the collection of positive contractions (of norm -strictly less than one) in a possibly non-unital C⋆-algebra form an approximate unit. The key step -is to show that this set is directed using the continuous functional calculus applied with the -functions `fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `fun x : ℝ≥0, x * (1 - x)⁻¹`, which are inverses on -the interval `{x : ℝ≥0 | x < 1}`. +This file shows that the collection of positive contractions (of norm strictly less than one) +in a possibly non-unital C⋆-algebra form a directed set. The key step uses the continuous functional +calculus applied with the functions `fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `fun x : ℝ≥0, x * (1 - x)⁻¹`, +which are inverses on the interval `{x : ℝ≥0 | x < 1}`. + +In addition, this file defines `IsIncreasingApproximateUnit` to be a filter `l` that is an +approximate unit contained in the closed unit ball of nonnegative elements. Every C⋆-algebra has +a filter generated by the sections `{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` and `‖a‖ < 1`, and +moreover, this filter is an increasing approximate unit. ## Main declarations @@ -22,11 +28,12 @@ the interval `{x : ℝ≥0 | x < 1}`. + `Set.InvOn.one_sub_one_add_inv`: the functions `f := fun x : ℝ≥0, 1 - (1 + x)⁻¹` and `g := fun x : ℝ≥0, x * (1 - x)⁻¹` are inverses on `{x : ℝ≥0 | x < 1}`. + `CStarAlgebra.directedOn_nonneg_ball`: the set `{x : A | 0 ≤ x} ∩ Metric.ball 0 1` is directed. - -## TODO - -+ Prove the approximate identity result by following Ken Davidson's proof in - "C*-Algebras by Example" ++ `Filter.IsIncreasingApproximateUnit`: a filter `l` is an *increasing approximate unit* if it is an + approximate unit contained in the closed unit ball of nonnegative elements. ++ `CStarAlgebra.approximateUnit`: the filter generated by the sections + `{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` with `‖a‖ < 1`. ++ `CStarAlgebra.increasingApproximateUnit`: the filter `CStarAlgebra.approximateUnit` is an + increasing approximate unit. -/ @@ -105,3 +112,195 @@ lemma CStarAlgebra.directedOn_nonneg_ball : have hab' : cfcₙ g a ≤ cfcₙ g a + cfcₙ g b := le_add_of_nonneg_right cfcₙ_nonneg_of_predicate exact CFC.monotoneOn_one_sub_one_add_inv cfcₙ_nonneg_of_predicate (cfcₙ_nonneg_of_predicate.trans hab') hab' + +section ApproximateUnit + +open Metric Filter Topology + +/-- An *increasing approximate unit* in a C⋆-algebra is an approximate unit contained in the +closed unit ball of nonnegative elements. -/ +structure Filter.IsIncreasingApproximateUnit (l : Filter A) extends l.IsApproximateUnit : Prop where + eventually_nonneg : ∀ᶠ x in l, 0 ≤ x + eventually_norm : ∀ᶠ x in l, ‖x‖ ≤ 1 + +namespace Filter.IsIncreasingApproximateUnit + +omit [StarOrderedRing A] in +lemma eventually_nnnorm {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, ‖x‖₊ ≤ 1 := + hl.eventually_norm + +lemma eventually_isSelfAdjoint {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, IsSelfAdjoint x := + hl.eventually_nonneg.mp <| .of_forall fun _ ↦ IsSelfAdjoint.of_nonneg + +lemma eventually_star_eq {l : Filter A} (hl : l.IsIncreasingApproximateUnit) : + ∀ᶠ x in l, star x = x := + hl.eventually_isSelfAdjoint.mp <| .of_forall fun _ ↦ IsSelfAdjoint.star_eq + +end Filter.IsIncreasingApproximateUnit + +namespace CStarAlgebra + +open Submodule in +/-- To show that `l` is a one-sided approximate unit for `A`, it suffices to verify it only for +`m : A` with `0 ≤ m` and `‖m‖ < 1`. -/ +lemma tendsto_mul_right_of_forall_nonneg_tendsto {l : Filter A} + (h : ∀ m, 0 ≤ m → ‖m‖ < 1 → Tendsto (· * m) l (𝓝 m)) (m : A) : + Tendsto (· * m) l (𝓝 m) := by + obtain ⟨n, c, x, rfl⟩ := mem_span_set'.mp <| by + show m ∈ span ℂ ({x | 0 ≤ x} ∩ ball 0 1) + simp [span_nonneg_inter_unitBall] + simp_rw [Finset.mul_sum] + refine tendsto_finset_sum _ fun i _ ↦ ?_ + simp_rw [mul_smul_comm] + exact tendsto_const_nhds.smul <| h (x i) (x i).2.1 <| by simpa using (x i).2.2 + +omit [PartialOrder A] in +/-- Multiplication on the left by `m` tends to `𝓝 m` if and only if multiplication on the right +does, provided the elements are eventually selfadjoint along the filter `l`. -/ +lemma tendsto_mul_left_iff_tendsto_mul_right {l : Filter A} (hl : ∀ᶠ x in l, IsSelfAdjoint x) : + (∀ m, Tendsto (m * ·) l (𝓝 m)) ↔ (∀ m, Tendsto (· * m) l (𝓝 m)) := by + refine ⟨fun h m ↦ ?_, fun h m ↦ ?_⟩ + all_goals + apply (star_star m ▸ (continuous_star.tendsto _ |>.comp <| h (star m))).congr' + filter_upwards [hl] with x hx + simp [hx.star_eq] + +variable (A) + +/-- The sections of positive strict contractions form a filter basis. -/ +lemma isBasis_nonneg_sections : + IsBasis (fun x : A ↦ 0 ≤ x ∧ ‖x‖ < 1) ({x | · ≤ x}) where + nonempty := ⟨0, by simp⟩ + inter {x y} hx hy := by + peel directedOn_nonneg_ball x (by simpa) y (by simpa) with z hz + exact ⟨by simpa using hz.1, fun a ha ↦ ⟨hz.2.1.trans ha, hz.2.2.trans ha⟩⟩ + +/-- The canonical approximate unit in a C⋆-algebra generated by the basis of sets +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a`. See also `CStarAlgebra.hasBasis_approximateUnit`. -/ +def approximateUnit : Filter A := + (isBasis_nonneg_sections A).filter ⊓ 𝓟 (closedBall 0 1) + +/-- The canonical approximate unit in a C⋆-algebra has a basis of sets +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a`. -/ +lemma hasBasis_approximateUnit : + (approximateUnit A).HasBasis (fun x : A ↦ 0 ≤ x ∧ ‖x‖ < 1) ({x | · ≤ x} ∩ closedBall 0 1) := + isBasis_nonneg_sections A |>.hasBasis.inf_principal (closedBall 0 1) + +/-- This is a common reasoning sequence in C⋆-algebra theory. If `0 ≤ x ≤ y ≤ 1`, then the norm +of `z - y * z` is controled by the norm of `star z * (1 - x) * z`, which is advantageous because the +latter is nonnegative. This is a key step in establishing the existence of an increasing approximate +unit in general C⋆-algebras. -/ +lemma nnnorm_sub_mul_self_le {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + {x y : A} (z : A) (hx₀ : 0 ≤ x) (hy : y ∈ Set.Icc x 1) {c : ℝ≥0} + (h : ‖star z * (1 - x) * z‖₊ ≤ c ^ 2) : + ‖z - y * z‖₊ ≤ c := by + nth_rw 1 [← one_mul z] + rw [← sqrt_sq c, le_sqrt_iff_sq_le, ← sub_mul, sq, ← CStarRing.nnnorm_star_mul_self] + simp only [star_mul, star_sub, star_one] + have hy₀ : y ∈ Set.Icc 0 1 := ⟨hx₀.trans hy.1, hy.2⟩ + have hy' : 1 - y ∈ Set.Icc 0 1 := Set.sub_mem_Icc_zero_iff_right.mpr hy₀ + rw [hy₀.1.star_eq, ← mul_assoc, mul_assoc (star _), ← sq] + refine nnnorm_le_nnnorm_of_nonneg_of_le (conjugate_nonneg (pow_nonneg hy'.1 2) _) ?_ |>.trans h + refine conjugate_le_conjugate ?_ _ + trans (1 - y) + · simpa using pow_antitone hy'.1 hy'.2 one_le_two + · gcongr + exact hy.1 + +/-- A variant of `nnnorm_sub_mul_self_le` which uses `‖·‖` instead of `‖·‖₊`. -/ +lemma norm_sub_mul_self_le {A : Type*} [CStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + {x y : A} (z : A) (hx₀ : 0 ≤ x) (hy : y ∈ Set.Icc x 1) + {c : ℝ} (hc : 0 ≤ c) (h : ‖star z * (1 - x) * z‖ ≤ c ^ 2) : + ‖z - y * z‖ ≤ c := + nnnorm_sub_mul_self_le z hx₀ hy h (c := ⟨c, hc⟩) + +variable {A} in +/-- A variant of `norm_sub_mul_self_le` for non-unital algebras that passes to the unitization. -/ +lemma norm_sub_mul_self_le_of_inr {x y : A} (z : A) (hx₀ : 0 ≤ x) (hxy : x ≤ y) (hy₁ : ‖y‖ ≤ 1) + {c : ℝ} (hc : 0 ≤ c) (h : ‖star (z : A⁺¹) * (1 - x) * z‖ ≤ c ^ 2) : + ‖z - y * z‖ ≤ c := by + rw [← norm_inr (𝕜 := ℂ), inr_sub, inr_mul] + refine norm_sub_mul_self_le _ ?_ ?_ hc h + · rwa [inr_nonneg_iff] + · have hy := hx₀.trans hxy + rw [Set.mem_Icc, inr_le_iff _ _ hx₀.isSelfAdjoint hy.isSelfAdjoint, + ← norm_le_one_iff_of_nonneg _, norm_inr] + exact ⟨hxy, hy₁⟩ + +variable {A} in +/-- This shows `CStarAlgebra.approximateUnit` is a one-sided approximate unit, but this is marked +`private` because it is only used to prove `CStarAlgebra.increasingApproximateUnit`. -/ +private lemma tendsto_mul_right_approximateUnit (m : A) : + Tendsto (· * m) (approximateUnit A) (𝓝 m) := by + refine tendsto_mul_right_of_forall_nonneg_tendsto (fun m hm₁ hm₂ ↦ ?_) m + rw [(hasBasis_approximateUnit A).tendsto_iff nhds_basis_closedBall] + intro ε hε + lift ε to ℝ≥0 using hε.le + rw [coe_pos] at hε + refine ⟨cfcₙ (fun y : ℝ≥0 ↦ 1 - (1 + y)⁻¹) (ε⁻¹ ^ 2 • m), + ⟨cfcₙ_nonneg_of_predicate, norm_cfcₙ_one_sub_one_add_inv_lt_one (ε⁻¹ ^ 2 • m)⟩, ?_⟩ + rintro x ⟨(hx₁ : _ ≤ x), hx₂⟩ + simp only [mem_closedBall, dist_eq_norm', zero_sub, norm_neg] at hx₂ ⊢ + rw [← coe_nnnorm, coe_le_coe] + have hx₀ : 0 ≤ x := cfcₙ_nonneg_of_predicate.trans hx₁ + rw [← inr_le_iff _ _ (.of_nonneg cfcₙ_nonneg_of_predicate) (.of_nonneg hx₀), + nnreal_cfcₙ_eq_cfc_inr _ _ (by simp [tsub_self]), inr_smul] at hx₁ + rw [← norm_inr (𝕜 := ℂ)] at hm₂ hx₂ + rw [← inr_nonneg_iff] at hx₀ hm₁ + rw [← nnnorm_inr (𝕜 := ℂ), inr_sub, inr_mul] + generalize (x : A⁺¹) = x, (m : A⁺¹) = m at * + set g : ℝ≥0 → ℝ≥0 := fun y ↦ 1 - (1 + y)⁻¹ + have hg : Continuous g := by + rw [continuous_iff_continuousOn_univ] + fun_prop (disch := intro _ _; positivity) + have hg' : ContinuousOn (fun y ↦ (1 + ε⁻¹ ^ 2 • y)⁻¹) (spectrum ℝ≥0 m) := + ContinuousOn.inv₀ (by fun_prop) fun _ _ ↦ by positivity + have hx : x ∈ Set.Icc 0 1 := mem_Icc_iff_norm_le_one.mpr ⟨hx₀, hx₂⟩ + have hx' : x ∈ Set.Icc _ 1 := ⟨hx₁, hx.2⟩ + refine nnnorm_sub_mul_self_le m cfc_nonneg_of_predicate hx' ?_ + suffices star m * (1 - cfc g (ε⁻¹ ^ 2 • m)) * m = + cfc (fun y : ℝ≥0 ↦ y * (1 + ε⁻¹ ^ 2 • y)⁻¹ * y) m by + rw [this] + refine nnnorm_cfc_nnreal_le fun y hy ↦ ?_ + field_simp + calc + y * ε ^ 2 * y / (ε ^ 2 + y) ≤ ε ^ 2 * 1 := by + rw [mul_div_assoc] + gcongr + · refine mul_le_of_le_one_left (zero_le _) ?_ + have hm' := hm₂.le + rw [norm_le_one_iff_of_nonneg m hm₁, ← cfc_id' ℝ≥0 m, ← cfc_one (R := ℝ≥0) m, + cfc_nnreal_le_iff _ _ _ (QuasispectrumRestricts.nnreal_of_nonneg hm₁)] at hm' + exact hm' y hy + · exact div_le_one (by positivity) |>.mpr le_add_self + _ = ε ^ 2 := mul_one _ + rw [cfc_mul _ _ m (continuousOn_id' _ |>.mul hg') (continuousOn_id' _), + cfc_mul _ _ m (continuousOn_id' _) hg', cfc_id' .., hm₁.star_eq] + congr + rw [← cfc_one (R := ℝ≥0) m, ← cfc_comp_smul _ _ _ hg.continuousOn hm₁, + ← cfc_tsub _ _ m (by simp [g]) hm₁ (by fun_prop) (Continuous.continuousOn <| by fun_prop)] + refine cfc_congr (fun y _ ↦ ?_) + simp [g, tsub_tsub_cancel_of_le] + +/-- The filter `CStarAlgebra.approximateUnit` generated by the sections +`{x | a ≤ x} ∩ closedBall 0 1` for `0 ≤ a` forms an increasing approximate unit. -/ +lemma increasingApproximateUnit : + IsIncreasingApproximateUnit (approximateUnit A) where + tendsto_mul_left := by + rw [tendsto_mul_left_iff_tendsto_mul_right] + · exact tendsto_mul_right_approximateUnit + · rw [(hasBasis_approximateUnit A).eventually_iff] + peel (hasBasis_approximateUnit A).ex_mem with x hx + exact ⟨hx, fun y hy ↦ (hx.1.trans hy.1).isSelfAdjoint⟩ + tendsto_mul_right := tendsto_mul_right_approximateUnit + eventually_nonneg := .filter_mono inf_le_left <| + (isBasis_nonneg_sections A).hasBasis.eventually_iff.mpr ⟨0, by simp⟩ + eventually_norm := .filter_mono inf_le_right <| by simp + neBot := hasBasis_approximateUnit A |>.neBot_iff.mpr + fun hx ↦ ⟨_, ⟨le_rfl, by simpa using hx.2.le⟩⟩ + +end CStarAlgebra + +end ApproximateUnit diff --git a/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean b/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean new file mode 100644 index 0000000000000..de68421d1db20 --- /dev/null +++ b/Mathlib/Analysis/CStarAlgebra/SpecialFunctions/PosPart.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2024 Jireh Loreaux. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jireh Loreaux +-/ +import Mathlib.Analysis.SpecialFunctions.ContinuousFunctionalCalculus.PosPart +import Mathlib.Analysis.CStarAlgebra.ContinuousFunctionalCalculus.Instances + +/-! # C⋆-algebraic facts about `a⁺` and `a⁻`. -/ + +variable {A : Type*} [NonUnitalCStarAlgebra A] [PartialOrder A] [StarOrderedRing A] + +namespace CStarAlgebra + +section SpanNonneg + +open Submodule + +/-- A C⋆-algebra is spanned by nonnegative elements of norm at most `r` -/ +lemma span_nonneg_inter_closedBall {r : ℝ} (hr : 0 < r) : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.closedBall 0 r) = ⊤ := by + rw [eq_top_iff, ← span_nonneg, span_le] + intro x hx + obtain (rfl | hx_pos) := eq_zero_or_norm_pos x + · exact zero_mem _ + · suffices (r * ‖x‖⁻¹ : ℂ)⁻¹ • ((r * ‖x‖⁻¹ : ℂ) • x) = x by + rw [← this] + refine smul_mem _ _ (subset_span <| Set.mem_inter ?_ ?_) + · norm_cast + exact smul_nonneg (by positivity) hx + · simp [mul_smul, norm_smul, abs_of_pos hr, inv_mul_cancel₀ hx_pos.ne'] + apply inv_smul_smul₀ + norm_cast + positivity + +/-- A C⋆-algebra is spanned by nonnegative elements of norm less than `r`. -/ +lemma span_nonneg_inter_ball {r : ℝ} (hr : 0 < r) : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.ball 0 r) = ⊤ := by + rw [eq_top_iff, ← span_nonneg_inter_closedBall (half_pos hr)] + gcongr + exact Metric.closedBall_subset_ball <| half_lt_self hr + +/-- A C⋆-algebra is spanned by nonnegative contractions. -/ +lemma span_nonneg_inter_unitClosedBall : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.closedBall 0 1) = ⊤ := + span_nonneg_inter_closedBall zero_lt_one + +/-- A C⋆-algebra is spanned by nonnegative strict contractions. -/ +lemma span_nonneg_inter_unitBall : + span ℂ ({x : A | 0 ≤ x} ∩ Metric.ball 0 1) = ⊤ := + span_nonneg_inter_ball zero_lt_one + +end SpanNonneg + +end CStarAlgebra diff --git a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean index 396d6e6043dca..959183c53b686 100644 --- a/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean +++ b/Mathlib/Analysis/SpecialFunctions/ContinuousFunctionalCalculus/PosPart.lean @@ -226,3 +226,29 @@ lemma posPart_negPart_unique {a b c : A} (habc : a = b - c) (hbc : b * c = 0) simp [Subtype.val_injective.extend_apply, f] end CFC + +section SpanNonneg + +variable {A : Type*} [NonUnitalRing A] [Module ℂ A] [SMulCommClass ℂ A A] [IsScalarTower ℂ A A] +variable [StarRing A] [TopologicalSpace A] [StarModule ℂ A] +variable [NonUnitalContinuousFunctionalCalculus ℝ (IsSelfAdjoint : A → Prop)] +variable [PartialOrder A] [StarOrderedRing A] + +open Submodule Complex +open scoped ComplexStarModule + +lemma CStarAlgebra.linear_combination_nonneg (x : A) : + ((ℜ x : A)⁺ - (ℜ x : A)⁻) + (I • (ℑ x : A)⁺ - I • (ℑ x : A)⁻) = x := by + rw [CFC.posPart_sub_negPart _ (ℜ x).2, ← smul_sub, CFC.posPart_sub_negPart _ (ℑ x).2, + realPart_add_I_smul_imaginaryPart x] + +/-- A C⋆-algebra is spanned by its nonnegative elements. -/ +lemma CStarAlgebra.span_nonneg : Submodule.span ℂ {a : A | 0 ≤ a} = ⊤ := by + refine eq_top_iff.mpr fun x _ => ?_ + rw [← CStarAlgebra.linear_combination_nonneg x] + apply_rules [sub_mem, Submodule.smul_mem, add_mem] + all_goals + refine subset_span ?_ + first | apply CFC.negPart_nonneg | apply CFC.posPart_nonneg + +end SpanNonneg From 8a7a0b7d493c44f352e65718307fce1c052d20e2 Mon Sep 17 00:00:00 2001 From: judithludwig Date: Wed, 27 Nov 2024 08:00:36 +0000 Subject: [PATCH 209/250] refactor(LinearAlgebra/Matrix/GeneralLinearGroup): Redefine `Matrix.SpecialLinearGroup.toGL`and remove duplicate `toLinear` (#19523) This PR redefines `Matrix.SpecialLinearGroup.toGL`, as the multiplicative map from the matrix group $$\mathrm{SL}_n(R)$$ to the matrix group $$\mathrm{GL}_n(R)$$, following a [Zulip discussion](https://leanprover.zulipchat.com/#narrow/channel/287929-mathlib4/topic/Matrix.2ESpecialLinearGroup.2EtoGL). It also removes a duplicate of `Matrix.GeneralLinearGroup.toLin`. Co-authored-by: judithludwig <150042763+judithludwig@users.noreply.github.com> --- .../Complex/UpperHalfPlane/Basic.lean | 6 ++-- .../Matrix/GeneralLinearGroup/Defs.lean | 31 +++++++++---------- .../Matrix/SpecialLinearGroup.lean | 9 ------ .../LinearAlgebra/Matrix/ToLinearEquiv.lean | 2 +- Mathlib/NumberTheory/Modular.lean | 6 ++-- 5 files changed, 22 insertions(+), 32 deletions(-) diff --git a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean index 6a3125c87d0c9..3a538cfd9d8ac 100644 --- a/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean +++ b/Mathlib/Analysis/Complex/UpperHalfPlane/Basic.lean @@ -364,12 +364,12 @@ end RealAddAction /- these next few lemmas are *not* flagged `@simp` because of the constructors on the RHS; instead we use the versions with coercions to `ℂ` as simp lemmas instead. -/ theorem modular_S_smul (z : ℍ) : ModularGroup.S • z = mk (-z : ℂ)⁻¹ z.im_inv_neg_coe_pos := by - rw [specialLinearGroup_apply]; simp [ModularGroup.S, neg_div, inv_neg, coeToGL] + rw [specialLinearGroup_apply]; simp [ModularGroup.S, neg_div, inv_neg, toGL] theorem modular_T_zpow_smul (z : ℍ) (n : ℤ) : ModularGroup.T ^ n • z = (n : ℝ) +ᵥ z := by rw [UpperHalfPlane.ext_iff, coe_vadd, add_comm, specialLinearGroup_apply, coe_mk] -- Porting note: added `coeToGL` and merged `rw` and `simp` - simp [coeToGL, ModularGroup.coe_T_zpow, + simp [toGL, ModularGroup.coe_T_zpow, of_apply, cons_val_zero, algebraMap.coe_one, Complex.ofReal_one, one_mul, cons_val_one, head_cons, algebraMap.coe_zero, zero_mul, zero_add, div_one] @@ -383,7 +383,7 @@ theorem exists_SL2_smul_eq_of_apply_zero_one_eq_zero (g : SL(2, ℝ)) (hc : g 1 ext1 ⟨z, hz⟩; ext1 suffices ↑a * z * a + b * a = b * a + a * a * z by -- Porting note: added `coeToGL` and merged `rw` and `simpa` - simpa [coeToGL, specialLinearGroup_apply, add_mul] + simpa [toGL, specialLinearGroup_apply, add_mul] ring theorem exists_SL2_smul_eq_of_apply_zero_one_ne_zero (g : SL(2, ℝ)) (hc : g 1 0 ≠ 0) : diff --git a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean index 62ce0b6d7abc0..6a59c9e32abf2 100644 --- a/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean +++ b/Mathlib/LinearAlgebra/Matrix/GeneralLinearGroup/Defs.lean @@ -5,6 +5,7 @@ Authors: Chris Birkbeck -/ import Mathlib.LinearAlgebra.Matrix.NonsingularInverse import Mathlib.LinearAlgebra.Matrix.SpecialLinearGroup +import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.Algebra.Ring.Subring.Units /-! @@ -69,7 +70,8 @@ def det : GL n R →* Rˣ where map_one' := Units.ext det_one map_mul' _ _ := Units.ext <| det_mul _ _ -/-- The `GL n R` and `Matrix.GeneralLinearGroup R n` groups are multiplicatively equivalent -/ +/-- The groups `GL n R` (notation for `Matrix.GeneralLinearGroup n R`) and +`LinearMap.GeneralLinearGroup R (n → R)` are multiplicatively equivalent -/ def toLin : GL n R ≃* LinearMap.GeneralLinearGroup R (n → R) := Units.mapEquiv toLinAlgEquiv'.toMulEquiv @@ -108,22 +110,19 @@ theorem coe_inv : ↑A⁻¹ = (↑A : Matrix n n R)⁻¹ := letI := A.invertible invOf_eq_nonsing_inv (↑A : Matrix n n R) -/-- An element of the matrix general linear group on `(n) [Fintype n]` can be considered as an -element of the endomorphism general linear group on `n → R`. -/ -def toLinear : GeneralLinearGroup n R ≃* LinearMap.GeneralLinearGroup R (n → R) := - Units.mapEquiv Matrix.toLinAlgEquiv'.toRingEquiv.toMulEquiv +@[deprecated (since := "2024-11-26")] alias toLinear := toLin -- Note that without the `@` and `‹_›`, Lean infers `fun a b ↦ _inst a b` instead of `_inst` as the -- decidability argument, which prevents `simp` from obtaining the instance by unification. -- These `fun a b ↦ _inst a b` terms also appear in the type of `A`, but simp doesn't get confused -- by them so for now we do not care. @[simp] -theorem coe_toLinear : (@toLinear n ‹_› ‹_› _ _ A : (n → R) →ₗ[R] n → R) = Matrix.mulVecLin A := +theorem coe_toLin : (@toLin n ‹_› ‹_› _ _ A : (n → R) →ₗ[R] n → R) = Matrix.mulVecLin A := rfl -- Porting note: is inserting toLinearEquiv here correct? @[simp] -theorem toLinear_apply (v : n → R) : (toLinear A).toLinearEquiv v = Matrix.mulVecLin (↑A) v := +theorem toLin_apply (v : n → R) : (toLin A).toLinearEquiv v = Matrix.mulVecLin (↑A) v := rfl end CoeLemmas @@ -195,17 +194,17 @@ namespace SpecialLinearGroup variable {n : Type u} [DecidableEq n] [Fintype n] {R : Type v} [CommRing R] --- Porting note: added implementation for the Coe -/-- The map from SL(n) to GL(n) underlying the coercion, forgetting the value of the determinant. --/ -@[coe] -def coeToGL (A : SpecialLinearGroup n R) : GL n R := - ⟨↑A, ↑A⁻¹, - congr_arg ((↑) : _ → Matrix n n R) (mul_inv_cancel A), - congr_arg ((↑) : _ → Matrix n n R) (inv_mul_cancel A)⟩ + +/-- `toGL` is the map from the special linear group to the general linear group. -/ +def toGL : Matrix.SpecialLinearGroup n R →* Matrix.GeneralLinearGroup n R where + toFun A := ⟨↑A, ↑A⁻¹, congr_arg (·.1) (mul_inv_cancel A), congr_arg (·.1) (inv_mul_cancel A)⟩ + map_one' := Units.ext rfl + map_mul' _ _ := Units.ext rfl + +@[deprecated (since := "2024-11-26")] alias coeToGL := toGL instance hasCoeToGeneralLinearGroup : Coe (SpecialLinearGroup n R) (GL n R) := - ⟨coeToGL⟩ + ⟨toGL⟩ @[simp] theorem coeToGL_det (g : SpecialLinearGroup n R) : diff --git a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean index be6c0ddf6440a..59bff4a8773bd 100644 --- a/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean +++ b/Mathlib/LinearAlgebra/Matrix/SpecialLinearGroup.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Anne Baanen. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Anne Baanen, Wen Yang -/ -import Mathlib.LinearAlgebra.GeneralLinearGroup import Mathlib.LinearAlgebra.Matrix.Adjugate import Mathlib.LinearAlgebra.Matrix.Transvection import Mathlib.RingTheory.RootsOfUnity.Basic @@ -211,14 +210,6 @@ theorem toLin'_injective : Function.Injective ↑(toLin' : SpecialLinearGroup n R →* (n → R) ≃ₗ[R] n → R) := fun _ _ h => Subtype.coe_injective <| Matrix.toLin'.injective <| LinearEquiv.toLinearMap_injective.eq_iff.mpr h -/-- `toGL` is the map from the special linear group to the general linear group -/ -def toGL : SpecialLinearGroup n R →* GeneralLinearGroup R (n → R) := - (GeneralLinearGroup.generalLinearEquiv _ _).symm.toMonoidHom.comp toLin' - --- Porting note (https://github.com/leanprover-community/mathlib4/issues/11036): broken dot notation -theorem coe_toGL (A : SpecialLinearGroup n R) : SpecialLinearGroup.toGL A = A.toLin'.toLinearMap := - rfl - variable {S : Type*} [CommRing S] /-- A ring homomorphism from `R` to `S` induces a group homomorphism from diff --git a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean index fefbfcaeaffce..c8f11ba565e7a 100644 --- a/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean +++ b/Mathlib/LinearAlgebra/Matrix/ToLinearEquiv.lean @@ -50,7 +50,7 @@ See `Matrix.toLinearEquiv` for the same map on arbitrary modules. -/ def toLinearEquiv' (P : Matrix n n R) (_ : Invertible P) : (n → R) ≃ₗ[R] n → R := GeneralLinearGroup.generalLinearEquiv _ _ <| - Matrix.GeneralLinearGroup.toLinear <| unitOfInvertible P + Matrix.GeneralLinearGroup.toLin <| unitOfInvertible P @[simp] theorem toLinearEquiv'_apply (P : Matrix n n R) (h : Invertible P) : diff --git a/Mathlib/NumberTheory/Modular.lean b/Mathlib/NumberTheory/Modular.lean index 59b083317b35b..acad10c21c22c 100644 --- a/Mathlib/NumberTheory/Modular.lean +++ b/Mathlib/NumberTheory/Modular.lean @@ -175,7 +175,7 @@ def lcRow0Extend {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : ![by refine LinearMap.GeneralLinearGroup.generalLinearEquiv ℝ (Fin 2 → ℝ) - (GeneralLinearGroup.toLinear (planeConformalMatrix (cd 0 : ℝ) (-(cd 1 : ℝ)) ?_)) + (GeneralLinearGroup.toLin (planeConformalMatrix (cd 0 : ℝ) (-(cd 1 : ℝ)) ?_)) norm_cast rw [neg_sq] exact hcd.sq_add_sq_ne_zero, LinearEquiv.refl ℝ (Fin 2 → ℝ)] @@ -210,12 +210,12 @@ theorem tendsto_lcRow0 {cd : Fin 2 → ℤ} (hcd : IsCoprime (cd 0) (cd 1)) : · simp only [Fin.isValue, Int.cast_one, map_apply_coe, RingHom.mapMatrix_apply, Int.coe_castRingHom, lcRow0_apply, map_apply, Fin.zero_eta, id_eq, Function.comp_apply, of_apply, cons_val', cons_val_zero, empty_val', cons_val_fin_one, lcRow0Extend_apply, - LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, + LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLin, val_planeConformalMatrix, neg_neg, mulVecLin_apply, mulVec, dotProduct, Fin.sum_univ_two, cons_val_one, head_cons, mB, f₁] · convert congr_arg (fun n : ℤ => (-n : ℝ)) g.det_coe.symm using 1 simp only [Fin.zero_eta, id_eq, Function.comp_apply, lcRow0Extend_apply, cons_val_zero, - LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLinear, + LinearMap.GeneralLinearGroup.coeFn_generalLinearEquiv, GeneralLinearGroup.coe_toLin, mulVecLin_apply, mulVec, dotProduct, det_fin_two, f₁] simp only [Fin.isValue, Fin.mk_one, val_planeConformalMatrix, neg_neg, of_apply, cons_val', empty_val', cons_val_fin_one, cons_val_one, head_fin_const, map_apply, Fin.sum_univ_two, From d51ea55ec492e5647536eea8573aed1bfe45680e Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Wed, 27 Nov 2024 08:35:36 +0000 Subject: [PATCH 210/250] chore: bump to nightly-2024-11-27 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index ca6a40f05dece..5d639a03e3bd2 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-26 +leanprover/lean4:nightly-2024-11-27 From 7e86316a6a2b02c4957950b7ee550b17506485f5 Mon Sep 17 00:00:00 2001 From: Jz Pan Date: Wed, 27 Nov 2024 08:55:06 +0000 Subject: [PATCH 211/250] feat(FieldTheory/LinearDisjoint): definition and basic properties of linearly disjoint of fields (#9651) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This PR adds basics about the linearly disjoint fields. Main definitions: - `IntermediateField.LinearDisjoint`: an intermediate field `A` of `E / F` and an abstract field `L` between `E / F` are linearly disjoint over `F`, if they are linearly disjoint as subalgebras (`Subalgebra.LinearDisjoint`). Main results: Equivalent characterization of linearly disjointness: - `IntermediateField.LinearDisjoint.linearIndependent_left`: if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains linearly independent over `L`. - `IntermediateField.LinearDisjoint.of_basis_left`: conversely, if there exists an `F`-basis of `A` which remains linearly independent over `L`, then `A` and `L` are linearly disjoint. - `IntermediateField.LinearDisjoint.linearIndependent_right`: if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains linearly independent over `A`. - `IntermediateField.LinearDisjoint.of_basis_right`: conversely, if there exists an `F`-basis of `L` which remains linearly independent over `A`, then `A` and `L` are linearly disjoint. - `IntermediateField.LinearDisjoint.linearIndependent_mul`: if `A` and `L` are linearly disjoint, then for any family of `F`-linearly independent elements `{ a_i }` of `A`, and any family of `F`-linearly independent elements `{ b_j }` of `L`, the family `{ a_i * b_j }` in `S` is also `F`-linearly independent. - `IntermediateField.LinearDisjoint.of_basis_mul`: conversely, if `{ a_i }` is an `F`-basis of `A`, if `{ b_j }` is an `F`-basis of `L`, such that the family `{ a_i * b_j }` in `E` is `F`-linearly independent, then `A` and `L` are linearly disjoint. Other main results: - `IntermediateField.LinearDisjoint.symm`, `IntermediateField.linearDisjoint_symm`: linearly disjoint is symmetric. - `IntermediateField.LinearDisjoint.rank_sup_of_isAlgebraic`, `IntermediateField.LinearDisjoint.finrank_sup`: if `A` and `B` are linearly disjoint, then the rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`. **TODO:** remove the algebraic assumptions (the proof becomes complicated). - `IntermediateField.LinearDisjoint.of_finrank_sup`: conversely, if `A` and `B` are finite extensions, such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, then `A` and `B` are linearly disjoint. - `IntermediateField.LinearDisjoint.of_finrank_coprime`: if the rank of `A` and `B` are coprime, then `A` and `B` are linearly disjoint. - `IntermediateField.LinearDisjoint.inf_eq_bot`: if `A` and `B` are linearly disjoint, then they are disjoint. --- Mathlib.lean | 1 + Mathlib/FieldTheory/LinearDisjoint.lean | 310 ++++++++++++++++++++++ Mathlib/LinearAlgebra/LinearDisjoint.lean | 26 +- Mathlib/RingTheory/LinearDisjoint.lean | 20 +- 4 files changed, 334 insertions(+), 23 deletions(-) create mode 100644 Mathlib/FieldTheory/LinearDisjoint.lean diff --git a/Mathlib.lean b/Mathlib.lean index 0195cb1242e40..26f7426ffef46 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2971,6 +2971,7 @@ import Mathlib.FieldTheory.JacobsonNoether import Mathlib.FieldTheory.KrullTopology import Mathlib.FieldTheory.KummerExtension import Mathlib.FieldTheory.Laurent +import Mathlib.FieldTheory.LinearDisjoint import Mathlib.FieldTheory.Minpoly.Basic import Mathlib.FieldTheory.Minpoly.Field import Mathlib.FieldTheory.Minpoly.IsConjRoot diff --git a/Mathlib/FieldTheory/LinearDisjoint.lean b/Mathlib/FieldTheory/LinearDisjoint.lean new file mode 100644 index 0000000000000..3ad86d6638858 --- /dev/null +++ b/Mathlib/FieldTheory/LinearDisjoint.lean @@ -0,0 +1,310 @@ +/- +Copyright (c) 2024 Jz Pan. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Jz Pan +-/ +import Mathlib.FieldTheory.Adjoin +import Mathlib.RingTheory.LinearDisjoint + +/-! + +# Linearly disjoint fields + +This file contains basics about the linearly disjoint fields. + +## Linear disjoint intermediate fields + +We adapt the definitions in . +See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. + +### Main definitions + +- `IntermediateField.LinearDisjoint`: an intermediate field `A` of `E / F` + and an abstract field `L` between `E / F` + (as a special case, two intermediate fields) are linearly disjoint over `F`, + if they are linearly disjoint as subalgebras (`Subalgebra.LinearDisjoint`). + +### Implementation notes + +The `Subalgebra.LinearDisjoint` is stated for two `Subalgebra`s. The original design of +`IntermediateField.LinearDisjoint` is also stated for two `IntermediateField`s +(see `IntermediateField.linearDisjoint_iff'` for the original statement). +But it's probably useful if one of them can be generalized to an abstract field +(see ). +This leads to the current design of `IntermediateField.LinearDisjoint` +which is for one `IntermediateField` and one abstract field. +It is not generalized to two abstract fields as this will break the dot notation. + +### Main results + +Equivalent characterization of linear disjointness: + +- `IntermediateField.LinearDisjoint.linearIndependent_left`: + if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains + linearly independent over `L`. + +- `IntermediateField.LinearDisjoint.of_basis_left`: + conversely, if there exists an `F`-basis of `A` which remains linearly independent over `L`, then + `A` and `L` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.linearIndependent_right`: + `IntermediateField.LinearDisjoint.linearIndependent_right'`: + if `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains + linearly independent over `A`. + +- `IntermediateField.LinearDisjoint.of_basis_right`: + `IntermediateField.LinearDisjoint.of_basis_right'`: + conversely, if there exists an `F`-basis of `L` which remains linearly independent over `A`, then + `A` and `L` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.linearIndependent_mul`: + `IntermediateField.LinearDisjoint.linearIndependent_mul'`: + if `A` and `L` are linearly disjoint, then for any family of + `F`-linearly independent elements `{ a_i }` of `A`, and any family of + `F`-linearly independent elements `{ b_j }` of `L`, the family `{ a_i * b_j }` in `S` is + also `F`-linearly independent. + +- `IntermediateField.LinearDisjoint.of_basis_mul`: + `IntermediateField.LinearDisjoint.of_basis_mul'`: + conversely, if `{ a_i }` is an `F`-basis of `A`, if `{ b_j }` is an `F`-basis of `L`, + such that the family `{ a_i * b_j }` in `E` is `F`-linearly independent, + then `A` and `L` are linearly disjoint. + +Other main results: + +- `IntermediateField.LinearDisjoint.symm`, `IntermediateField.linearDisjoint_comm`: + linear disjointness is symmetric. + +- `IntermediateField.LinearDisjoint.rank_sup_of_isAlgebraic`, + `IntermediateField.LinearDisjoint.finrank_sup`: + if `A` and `B` are linearly disjoint, + then the rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`. + + **TODO:** remove the algebraic assumptions (the proof becomes complicated). + +- `IntermediateField.LinearDisjoint.of_finrank_sup`: + conversely, if `A` and `B` are finite extensions, + such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, + then `A` and `B` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.of_finrank_coprime`: + if the rank of `A` and `B` are coprime, + then `A` and `B` are linearly disjoint. + +- `IntermediateField.LinearDisjoint.inf_eq_bot`: + if `A` and `B` are linearly disjoint, then they are disjoint. + +## Tags + +linearly disjoint, linearly independent, tensor product + +-/ + +open scoped TensorProduct + +open Module IntermediateField + +noncomputable section + +universe u v w + +namespace IntermediateField + +variable {F : Type u} {E : Type v} [Field F] [Field E] [Algebra F E] + +variable (A B : IntermediateField F E) + +variable (L : Type w) [Field L] [Algebra F L] [Algebra L E] [IsScalarTower F L E] + +/-- If `A` is an intermediate field of `E / F`, and `E / L / F` is a field extension tower, +then `A` and `L` are linearly disjoint, if they are linearly disjoint as subalgebras of `E` +(`Subalgebra.LinearDisjoint`). -/ +protected abbrev LinearDisjoint : Prop := + A.toSubalgebra.LinearDisjoint (IsScalarTower.toAlgHom F L E).range + +theorem linearDisjoint_iff : + A.LinearDisjoint L ↔ A.toSubalgebra.LinearDisjoint (IsScalarTower.toAlgHom F L E).range := + Iff.rfl + +variable {A B L} + +/-- Two intermediate fields are linearly disjoint if and only if +they are linearly disjoint as subalgebras. -/ +theorem linearDisjoint_iff' : + A.LinearDisjoint B ↔ A.toSubalgebra.LinearDisjoint B.toSubalgebra := by + rw [linearDisjoint_iff] + congr! + ext; simp + +/-- Linear disjointness is symmetric. -/ +theorem LinearDisjoint.symm (H : A.LinearDisjoint B) : B.LinearDisjoint A := + linearDisjoint_iff'.2 (linearDisjoint_iff'.1 H).symm + +/-- Linear disjointness is symmetric. -/ +theorem linearDisjoint_comm : A.LinearDisjoint B ↔ B.LinearDisjoint A := + ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ + +namespace LinearDisjoint + +variable (A) in +theorem self_right : A.LinearDisjoint F := Subalgebra.LinearDisjoint.bot_right _ + +variable (A) in +theorem bot_right : A.LinearDisjoint (⊥ : IntermediateField F E) := + linearDisjoint_iff'.2 (Subalgebra.LinearDisjoint.bot_right _) + +variable (F E L) in +theorem bot_left : (⊥ : IntermediateField F E).LinearDisjoint L := + Subalgebra.LinearDisjoint.bot_left _ + +/-- If `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `A` remains +linearly independent over `L`. -/ +theorem linearIndependent_left (H : A.LinearDisjoint L) + {ι : Type*} {a : ι → A} (ha : LinearIndependent F a) : LinearIndependent L (A.val ∘ a) := + (Subalgebra.LinearDisjoint.linearIndependent_left_of_flat H ha).map_of_injective_injective + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)) (AddMonoidHom.id E) + (by simp) (by simp) (fun _ _ ↦ by simp_rw [Algebra.smul_def]; rfl) + +/-- If there exists an `F`-basis of `A` which remains linearly independent over `L`, then +`A` and `L` are linearly disjoint. -/ +theorem of_basis_left {ι : Type*} (a : Basis ι F A) + (H : LinearIndependent L (A.val ∘ a)) : A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_left _ _ a <| H.map_of_surjective_injective + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)) (AddMonoidHom.id E) + (AlgEquiv.surjective _) (by simp) (fun _ _ ↦ by simp_rw [Algebra.smul_def]; rfl) + +/-- If `A` and `B` are linearly disjoint, then any `F`-linearly independent family on `B` remains +linearly independent over `A`. -/ +theorem linearIndependent_right (H : A.LinearDisjoint B) + {ι : Type*} {b : ι → B} (hb : LinearIndependent F b) : LinearIndependent A (B.val ∘ b) := + (linearDisjoint_iff'.1 H).linearIndependent_right_of_flat hb + +/-- If there exists an `F`-basis of `B` which remains linearly independent over `A`, then +`A` and `B` are linearly disjoint. -/ +theorem of_basis_right {ι : Type*} (b : Basis ι F B) + (H : LinearIndependent A (B.val ∘ b)) : A.LinearDisjoint B := + linearDisjoint_iff'.2 (.of_basis_right _ _ b H) + +/-- If `A` and `L` are linearly disjoint, then any `F`-linearly independent family on `L` remains +linearly independent over `A`. -/ +theorem linearIndependent_right' (H : A.LinearDisjoint L) {ι : Type*} {b : ι → L} + (hb : LinearIndependent F b) : LinearIndependent A (algebraMap L E ∘ b) := by + apply Subalgebra.LinearDisjoint.linearIndependent_right_of_flat H <| hb.map' _ + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.ker + +/-- If there exists an `F`-basis of `L` which remains linearly independent over `A`, then +`A` and `L` are linearly disjoint. -/ +theorem of_basis_right' {ι : Type*} (b : Basis ι F L) + (H : LinearIndependent A (algebraMap L E ∘ b)) : A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_right _ _ + (b.map (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv) H + +/-- If `A` and `B` are linearly disjoint, then for any `F`-linearly independent families +`{ u_i }`, `{ v_j }` of `A`, `B`, the products `{ u_i * v_j }` +are linearly independent over `F`. -/ +theorem linearIndependent_mul (H : A.LinearDisjoint B) {κ ι : Type*} {a : κ → A} {b : ι → B} + (ha : LinearIndependent F a) (hb : LinearIndependent F b) : + LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * (b i.2).1 := + (linearDisjoint_iff'.1 H).linearIndependent_mul_of_flat_left ha hb + +/-- If `A` and `L` are linearly disjoint, then for any `F`-linearly independent families +`{ u_i }`, `{ v_j }` of `A`, `L`, the products `{ u_i * v_j }` +are linearly independent over `F`. -/ +theorem linearIndependent_mul' (H : A.LinearDisjoint L) {κ ι : Type*} {a : κ → A} {b : ι → L} + (ha : LinearIndependent F a) (hb : LinearIndependent F b) : + LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * algebraMap L E (b i.2) := by + apply Subalgebra.LinearDisjoint.linearIndependent_mul_of_flat_left H ha <| hb.map' _ + (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.ker + +/-- If there are `F`-bases `{ u_i }`, `{ v_j }` of `A`, `B`, such that the products +`{ u_i * v_j }` are linearly independent over `F`, then `A` and `B` are linearly disjoint. -/ +theorem of_basis_mul {κ ι : Type*} (a : Basis κ F A) (b : Basis ι F B) + (H : LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * (b i.2).1) : A.LinearDisjoint B := + linearDisjoint_iff'.2 (.of_basis_mul _ _ a b H) + +/-- If there are `F`-bases `{ u_i }`, `{ v_j }` of `A`, `L`, such that the products +`{ u_i * v_j }` are linearly independent over `F`, then `A` and `L` are linearly disjoint. -/ +theorem of_basis_mul' {κ ι : Type*} (a : Basis κ F A) (b : Basis ι F L) + (H : LinearIndependent F fun (i : κ × ι) ↦ (a i.1).1 * algebraMap L E (b i.2)) : + A.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_basis_mul _ _ a + (b.map (AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv) H + +theorem of_le_left {A' : IntermediateField F E} (H : A.LinearDisjoint L) + (h : A' ≤ A) : A'.LinearDisjoint L := + Subalgebra.LinearDisjoint.of_le_left_of_flat H h + +theorem of_le_right {B' : IntermediateField F E} (H : A.LinearDisjoint B) + (h : B' ≤ B) : A.LinearDisjoint B' := + linearDisjoint_iff'.2 ((linearDisjoint_iff'.1 H).of_le_right_of_flat h) + +/-- Similar to `IntermediateField.LinearDisjoint.of_le_right` but this is for abstract fields. -/ +theorem of_le_right' (H : A.LinearDisjoint L) (L' : Type*) [Field L'] + [Algebra F L'] [Algebra L' L] [IsScalarTower F L' L] + [Algebra L' E] [IsScalarTower F L' E] [IsScalarTower L' L E] : A.LinearDisjoint L' := by + refine Subalgebra.LinearDisjoint.of_le_right_of_flat H ?_ + convert AlgHom.range_comp_le_range (IsScalarTower.toAlgHom F L' L) (IsScalarTower.toAlgHom F L E) + ext; exact IsScalarTower.algebraMap_apply L' L E _ + +/-- If `A` and `B` are linearly disjoint, `A'` and `B'` are contained in `A` and `B`, +respectively, then `A'` and `B'` are also linearly disjoint. -/ +theorem of_le {A' B' : IntermediateField F E} (H : A.LinearDisjoint B) + (hA : A' ≤ A) (hB : B' ≤ B) : A'.LinearDisjoint B' := + H.of_le_left hA |>.of_le_right hB + +/-- Similar to `IntermediateField.LinearDisjoint.of_le` but this is for abstract fields. -/ +theorem of_le' {A' : IntermediateField F E} (H : A.LinearDisjoint L) + (hA : A' ≤ A) (L' : Type*) [Field L'] + [Algebra F L'] [Algebra L' L] [IsScalarTower F L' L] + [Algebra L' E] [IsScalarTower F L' E] [IsScalarTower L' L E] : A'.LinearDisjoint L' := + H.of_le_left hA |>.of_le_right' L' + +/-- If `A` and `B` are linearly disjoint over `F`, then their intersection is equal to `F`. -/ +theorem inf_eq_bot (H : A.LinearDisjoint B) : + A ⊓ B = ⊥ := toSubalgebra_injective (linearDisjoint_iff'.1 H).inf_eq_bot + +/-- If `A` and `A` itself are linearly disjoint over `F`, then it is equal to `F`. -/ +theorem eq_bot_of_self (H : A.LinearDisjoint A) : A = ⊥ := + inf_idem A ▸ H.inf_eq_bot + +/-- If `A` and `B` are linearly disjoint over `F`, and at least one them are algebraic, then the +rank of `A ⊔ B` is equal to the product of that of `A` and `B`. Note that this result is +also true without algebraic assumption, but the proof becomes very complicated. -/ +theorem rank_sup_of_isAlgebraic (H : A.LinearDisjoint B) + (halg : Algebra.IsAlgebraic F A ∨ Algebra.IsAlgebraic F B) : + Module.rank F ↥(A ⊔ B) = Module.rank F A * Module.rank F B := + have h := le_sup_toSubalgebra A B + (rank_sup_le_of_isAlgebraic A B halg).antisymm <| + (linearDisjoint_iff'.1 H).rank_sup_of_free.ge.trans <| + (Subalgebra.inclusion h).toLinearMap.rank_le_of_injective (Subalgebra.inclusion_injective h) + +/-- If `A` and `B` are linearly disjoint over `F`, then the `Module.finrank` of +`A ⊔ B` is equal to the product of that of `A` and `B`. -/ +theorem finrank_sup (H : A.LinearDisjoint B) : finrank F ↥(A ⊔ B) = finrank F A * finrank F B := by + by_cases h : FiniteDimensional F A + · simpa only [map_mul] using + congr(Cardinal.toNat $(H.rank_sup_of_isAlgebraic (.inl inferInstance))) + rw [FiniteDimensional, ← rank_lt_aleph0_iff, not_lt] at h + have := LinearMap.rank_le_of_injective _ <| Submodule.inclusion_injective <| + show Subalgebra.toSubmodule A.toSubalgebra ≤ Subalgebra.toSubmodule (A ⊔ B).toSubalgebra by simp + rw [show finrank F A = 0 from Cardinal.toNat_apply_of_aleph0_le h, + show finrank F ↥(A ⊔ B) = 0 from Cardinal.toNat_apply_of_aleph0_le (h.trans this), zero_mul] + +/-- If `A` and `B` are finite extensions of `F`, +such that rank of `A ⊔ B` is equal to the product of the rank of `A` and `B`, +then `A` and `B` are linearly disjoint. -/ +theorem of_finrank_sup [FiniteDimensional F A] [FiniteDimensional F B] + (H : finrank F ↥(A ⊔ B) = finrank F A * finrank F B) : A.LinearDisjoint B := + linearDisjoint_iff'.2 <| .of_finrank_sup_of_free (by rwa [← sup_toSubalgebra_of_left]) + +/-- If `A` and `L` have coprime degree over `F`, then they are linearly disjoint. -/ +theorem of_finrank_coprime (H : (finrank F A).Coprime (finrank F L)) : A.LinearDisjoint L := + letI : Field (AlgHom.range (IsScalarTower.toAlgHom F L E)) := + inferInstanceAs <| Field (AlgHom.fieldRange (IsScalarTower.toAlgHom F L E)) + letI : Field A.toSubalgebra := inferInstanceAs <| Field A + Subalgebra.LinearDisjoint.of_finrank_coprime_of_free <| by + rwa [(AlgEquiv.ofInjectiveField (IsScalarTower.toAlgHom F L E)).toLinearEquiv.finrank_eq] at H + +end LinearDisjoint + +end IntermediateField diff --git a/Mathlib/LinearAlgebra/LinearDisjoint.lean b/Mathlib/LinearAlgebra/LinearDisjoint.lean index 18ed93e9d07ed..260ea22120c2e 100644 --- a/Mathlib/LinearAlgebra/LinearDisjoint.lean +++ b/Mathlib/LinearAlgebra/LinearDisjoint.lean @@ -25,7 +25,7 @@ Let `M` and `N` be `R`-submodules in `S` (`Submodule R S`). `M.LinearDisjoint N`), if the natural `R`-linear map `M ⊗[R] N →ₗ[R] S` (`Submodule.mulMap M N`) induced by the multiplication in `S` is injective. -The following is the first equivalent characterization of linearly disjointness: +The following is the first equivalent characterization of linear disjointness: - `Submodule.LinearDisjoint.linearIndependent_left_of_flat`: if `M` and `N` are linearly disjoint, if `N` is a flat `R`-module, then for any family of @@ -49,7 +49,7 @@ Dually, we have: conversely, if `{ n_i }` is an `R`-basis of `N`, which is also `M`-linearly independent, then `M` and `N` are linearly disjoint. -The following is the second equivalent characterization of linearly disjointness: +The following is the second equivalent characterization of linear disjointness: - `Submodule.LinearDisjoint.linearIndependent_mul_of_flat`: if `M` and `N` are linearly disjoint, if one of `M` and `N` is flat, then for any family of @@ -64,15 +64,15 @@ The following is the second equivalent characterization of linearly disjointness ## Other main results -- `Submodule.LinearDisjoint.symm_of_commute`, `Submodule.linearDisjoint_symm_of_commute`: - linearly disjoint is symmetric under some commutative conditions. +- `Submodule.LinearDisjoint.symm_of_commute`, `Submodule.linearDisjoint_comm_of_commute`: + linear disjointness is symmetric under some commutative conditions. - `Submodule.linearDisjoint_op`: - linearly disjoint is preserved by taking multiplicative opposite. + linear disjointness is preserved by taking multiplicative opposite. - `Submodule.LinearDisjoint.of_le_left_of_flat`, `Submodule.LinearDisjoint.of_le_right_of_flat`, `Submodule.LinearDisjoint.of_le_of_flat_left`, `Submodule.LinearDisjoint.of_le_of_flat_right`: - linearly disjoint is preserved by taking submodules under some flatness conditions. + linear disjointness is preserved by taking submodules under some flatness conditions. - `Submodule.LinearDisjoint.of_linearDisjoint_fg_left`, `Submodule.LinearDisjoint.of_linearDisjoint_fg_right`, @@ -162,7 +162,7 @@ theorem LinearDisjoint.of_subsingleton [Subsingleton R] : M.LinearDisjoint N := haveI : Subsingleton S := Module.subsingleton R S exact ⟨Function.injective_of_subsingleton _⟩ -/-- Linearly disjoint is preserved by taking multiplicative opposite. -/ +/-- Linear disjointness is preserved by taking multiplicative opposite. -/ theorem linearDisjoint_op : M.LinearDisjoint N ↔ (equivOpposite.symm (MulOpposite.op N)).LinearDisjoint (equivOpposite.symm (MulOpposite.op M)) := by @@ -171,14 +171,14 @@ theorem linearDisjoint_op : alias ⟨LinearDisjoint.op, LinearDisjoint.of_op⟩ := linearDisjoint_op -/-- Linearly disjoint is symmetric if elements in the module commute. -/ +/-- Linear disjointness is symmetric if elements in the module commute. -/ theorem LinearDisjoint.symm_of_commute (H : M.LinearDisjoint N) (hc : ∀ (m : M) (n : N), Commute m.1 n.1) : N.LinearDisjoint M := by rw [linearDisjoint_iff, mulMap_comm_of_commute M N hc] exact ((TensorProduct.comm R N M).toEquiv.injective_comp _).2 H.injective -/-- Linearly disjoint is symmetric if elements in the module commute. -/ -theorem linearDisjoint_symm_of_commute +/-- Linear disjointness is symmetric if elements in the module commute. -/ +theorem linearDisjoint_comm_of_commute (hc : ∀ (m : M) (n : N), Commute m.1 n.1) : M.LinearDisjoint N ↔ N.LinearDisjoint M := ⟨fun H ↦ H.symm_of_commute hc, fun H ↦ H.symm_of_commute fun _ _ ↦ (hc _ _).symm⟩ @@ -284,12 +284,12 @@ variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable {M N : Submodule R S} -/-- Linearly disjoint is symmetric in a commutative ring. -/ +/-- Linear disjointness is symmetric in a commutative ring. -/ theorem LinearDisjoint.symm (H : M.LinearDisjoint N) : N.LinearDisjoint M := H.symm_of_commute fun _ _ ↦ mul_comm _ _ -/-- Linearly disjoint is symmetric in a commutative ring. -/ -theorem linearDisjoint_symm : M.LinearDisjoint N ↔ N.LinearDisjoint M := +/-- Linear disjointness is symmetric in a commutative ring. -/ +theorem linearDisjoint_comm : M.LinearDisjoint N ↔ N.LinearDisjoint M := ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ end CommSemiring diff --git a/Mathlib/RingTheory/LinearDisjoint.lean b/Mathlib/RingTheory/LinearDisjoint.lean index 11a726e495c0b..b03dac78930c7 100644 --- a/Mathlib/RingTheory/LinearDisjoint.lean +++ b/Mathlib/RingTheory/LinearDisjoint.lean @@ -32,7 +32,7 @@ See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. ## Main results -### Equivalent characterization of linearly disjointness +### Equivalent characterization of linear disjointness - `Subalgebra.LinearDisjoint.linearIndependent_left_of_flat`: if `A` and `B` are linearly disjoint, and if `B` is a flat `R`-module, then for any family of @@ -63,8 +63,8 @@ See the file `Mathlib/LinearAlgebra/LinearDisjoint.lean` for details. ### Other main results -- `Subalgebra.LinearDisjoint.symm_of_commute`, `Subalgebra.linearDisjoint_symm_of_commute`: - linearly disjoint is symmetric under some commutative conditions. +- `Subalgebra.LinearDisjoint.symm_of_commute`, `Subalgebra.linearDisjoint_comm_of_commute`: + linear disjointness is symmetric under some commutative conditions. - `Subalgebra.LinearDisjoint.bot_left`, `Subalgebra.LinearDisjoint.bot_right`: the image of `R` in `S` is linearly disjoint with any other subalgebras. @@ -104,7 +104,7 @@ linearly disjoint, linearly independent, tensor product -/ -open scoped Classical TensorProduct +open scoped TensorProduct noncomputable section @@ -133,13 +133,13 @@ variable {A B} theorem LinearDisjoint.of_subsingleton [Subsingleton R] : A.LinearDisjoint B := Submodule.LinearDisjoint.of_subsingleton -/-- Linearly disjoint is symmetric if elements in the module commute. -/ +/-- Linear disjointness is symmetric if elements in the module commute. -/ theorem LinearDisjoint.symm_of_commute (H : A.LinearDisjoint B) (hc : ∀ (a : A) (b : B), Commute a.1 b.1) : B.LinearDisjoint A := Submodule.LinearDisjoint.symm_of_commute H hc -/-- Linearly disjoint is symmetric if elements in the module commute. -/ -theorem linearDisjoint_symm_of_commute +/-- Linear disjointness is symmetric if elements in the module commute. -/ +theorem linearDisjoint_comm_of_commute (hc : ∀ (a : A) (b : B), Commute a.1 b.1) : A.LinearDisjoint B ↔ B.LinearDisjoint A := ⟨fun H ↦ H.symm_of_commute hc, fun H ↦ H.symm_of_commute fun _ _ ↦ (hc _ _).symm⟩ @@ -167,12 +167,12 @@ variable [CommSemiring R] [CommSemiring S] [Algebra R S] variable {A B : Subalgebra R S} -/-- Linearly disjoint is symmetric in a commutative ring. -/ +/-- Linear disjointness is symmetric in a commutative ring. -/ theorem LinearDisjoint.symm (H : A.LinearDisjoint B) : B.LinearDisjoint A := H.symm_of_commute fun _ _ ↦ mul_comm _ _ -/-- Linearly disjoint is symmetric in a commutative ring. -/ -theorem linearDisjoint_symm : A.LinearDisjoint B ↔ B.LinearDisjoint A := +/-- Linear disjointness is symmetric in a commutative ring. -/ +theorem linearDisjoint_comm : A.LinearDisjoint B ↔ B.LinearDisjoint A := ⟨LinearDisjoint.symm, LinearDisjoint.symm⟩ namespace LinearDisjoint From f899f09c0d59f45ba460b323c102500e38b34f26 Mon Sep 17 00:00:00 2001 From: sgouezel Date: Wed, 27 Nov 2024 09:04:45 +0000 Subject: [PATCH 212/250] chore: avoid importing `ContDiff.Defs` in `FDeriv.Analytic` (#19374) For this, move the results needing `ContDiff` to two new files. The reason of this change is that I will import `FDeriv.Analytic` in `ContDiff.Defs` in #17152 --- Mathlib.lean | 2 + Mathlib/Analysis/Analytic/Within.lean | 29 +++++-- .../Analysis/Calculus/ContDiff/Analytic.lean | 55 +++++++++++++ .../Calculus/ContDiff/CPolynomial.lean | 63 +++++++++++++++ .../Analysis/Calculus/FDeriv/Analytic.lean | 79 +++---------------- .../Fourier/FourierTransformDeriv.lean | 3 +- .../Analysis/SpecialFunctions/ExpDeriv.lean | 2 + .../Geometry/Manifold/AnalyticManifold.lean | 2 +- 8 files changed, 155 insertions(+), 80 deletions(-) create mode 100644 Mathlib/Analysis/Calculus/ContDiff/Analytic.lean create mode 100644 Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean diff --git a/Mathlib.lean b/Mathlib.lean index 26f7426ffef46..21358bdd55951 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1120,8 +1120,10 @@ import Mathlib.Analysis.Calculus.BumpFunction.InnerProduct import Mathlib.Analysis.Calculus.BumpFunction.Normed import Mathlib.Analysis.Calculus.Conformal.InnerProduct import Mathlib.Analysis.Calculus.Conformal.NormedSpace +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.Basic import Mathlib.Analysis.Calculus.ContDiff.Bounds +import Mathlib.Analysis.Calculus.ContDiff.CPolynomial import Mathlib.Analysis.Calculus.ContDiff.Defs import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.ContDiff.FaaDiBruno diff --git a/Mathlib/Analysis/Analytic/Within.lean b/Mathlib/Analysis/Analytic/Within.lean index 67e0db70351e1..fd9bd854f032d 100644 --- a/Mathlib/Analysis/Analytic/Within.lean +++ b/Mathlib/Analysis/Analytic/Within.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions +import Mathlib.Analysis.Analytic.ChangeOrigin /-! # Properties of analyticity restricted to a set @@ -40,8 +41,8 @@ lemma analyticWithinAt_of_singleton_mem {f : E → F} {s : Set E} {x : E} (h : { AnalyticWithinAt 𝕜 f s x := by rcases mem_nhdsWithin.mp h with ⟨t, ot, xt, st⟩ rcases Metric.mem_nhds_iff.mp (ot.mem_nhds xt) with ⟨r, r0, rt⟩ - exact ⟨constFormalMultilinearSeries 𝕜 E (f x), .ofReal r, { - r_le := by simp only [FormalMultilinearSeries.constFormalMultilinearSeries_radius, le_top] + exact ⟨constFormalMultilinearSeries 𝕜 E (f x), .ofReal r, + { r_le := by simp only [FormalMultilinearSeries.constFormalMultilinearSeries_radius, le_top] r_pos := by positivity hasSum := by intro y ys yr @@ -63,10 +64,10 @@ lemma analyticOn_of_locally_analyticOn {f : E → F} {s : Set E} rcases h x m with ⟨u, ou, xu, fu⟩ rcases Metric.mem_nhds_iff.mp (ou.mem_nhds xu) with ⟨r, r0, ru⟩ rcases fu x ⟨m, xu⟩ with ⟨p, t, fp⟩ - exact ⟨p, min (.ofReal r) t, { - r_pos := lt_min (by positivity) fp.r_pos - r_le := min_le_of_right_le fp.r_le - hasSum := by + exact ⟨p, min (.ofReal r) t, + { r_pos := lt_min (by positivity) fp.r_pos + r_le := min_le_of_right_le fp.r_le + hasSum := by intro y ys yr simp only [EMetric.mem_ball, lt_min_iff, edist_lt_ofReal, dist_zero_right] at yr apply fp.hasSum @@ -88,8 +89,8 @@ lemma IsOpen.analyticOn_iff_analyticOnNhd {f : E → F} {s : Set E} (hs : IsOpen intro hf x m rcases Metric.mem_nhds_iff.mp (hs.mem_nhds m) with ⟨r, r0, rs⟩ rcases hf x m with ⟨p, t, fp⟩ - exact ⟨p, min (.ofReal r) t, { - r_pos := lt_min (by positivity) fp.r_pos + exact ⟨p, min (.ofReal r) t, + { r_pos := lt_min (by positivity) fp.r_pos r_le := min_le_of_right_le fp.r_le hasSum := by intro y ym @@ -200,3 +201,15 @@ lemma analyticWithinAt_iff_exists_analyticAt' [CompleteSpace F] {f : E → F} {s exact ⟨g, by filter_upwards [self_mem_nhdsWithin] using hf, hg⟩ alias ⟨AnalyticWithinAt.exists_analyticAt, _⟩ := analyticWithinAt_iff_exists_analyticAt' + +lemma AnalyticWithinAt.exists_mem_nhdsWithin_analyticOn + [CompleteSpace F] {f : E → F} {s : Set E} {x : E} (h : AnalyticWithinAt 𝕜 f s x) : + ∃ u ∈ 𝓝[insert x s] x, AnalyticOn 𝕜 f u := by + obtain ⟨g, -, h'g, hg⟩ : ∃ g, f x = g x ∧ EqOn f g (insert x s) ∧ AnalyticAt 𝕜 g x := + h.exists_analyticAt + let u := insert x s ∩ {y | AnalyticAt 𝕜 g y} + refine ⟨u, ?_, ?_⟩ + · exact inter_mem_nhdsWithin _ ((isOpen_analyticAt 𝕜 g).mem_nhds hg) + · intro y hy + have : AnalyticWithinAt 𝕜 g u y := hy.2.analyticWithinAt + exact this.congr (h'g.mono (inter_subset_left)) (h'g (inter_subset_left hy)) diff --git a/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean b/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean new file mode 100644 index 0000000000000..8ba3cda66b843 --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/Analytic.lean @@ -0,0 +1,55 @@ +/- +Copyright (c) 2021 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Analysis.Calculus.ContDiff.Defs +import Mathlib.Analysis.Calculus.FDeriv.Analytic + +/-! +# Analytic functions are `C^∞`. +-/ + +open Set ContDiff + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] + {E : Type*} [NormedAddCommGroup E] [NormedSpace 𝕜 E] + {F : Type*} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + {f : E → F} {s : Set E} {x : E} {n : WithTop ℕ∞} + +/-- An analytic function is infinitely differentiable. -/ +protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) : + ContDiffOn 𝕜 n f s := by + let t := { x | AnalyticAt 𝕜 f x } + suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] + have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx + have t_open : IsOpen t := isOpen_analyticAt 𝕜 f + exact contDiffOn_of_continuousOn_differentiableOn + (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + (fun m _ ↦ (H.iteratedFDeriv m).differentiableOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + +/-- An analytic function on the whole space is infinitely differentiable there. -/ +theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) : + ContDiff 𝕜 n f := by + rw [← contDiffOn_univ] + exact h.contDiffOn + +theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : ℕ∞} : + ContDiffAt 𝕜 n f x := by + obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd + exact hf.contDiffOn.contDiffAt hs + +protected lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} + (h : AnalyticWithinAt 𝕜 f s x) {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x := by + rcases h.exists_analyticAt with ⟨g, fx, fg, hg⟩ + exact hg.contDiffAt.contDiffWithinAt.congr (fg.mono (subset_insert _ _)) fx + +protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E} + (h : AnalyticOn 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s := + fun x m ↦ (h x m).contDiffWithinAt + +@[deprecated (since := "2024-09-26")] +alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn diff --git a/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean new file mode 100644 index 0000000000000..3a66c7c3d570a --- /dev/null +++ b/Mathlib/Analysis/Calculus/ContDiff/CPolynomial.lean @@ -0,0 +1,63 @@ +/- +Copyright (c) 2021 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yury Kudryashov +-/ +import Mathlib.Analysis.Calculus.FDeriv.Analytic +import Mathlib.Analysis.Calculus.ContDiff.Defs + +/-! +# Higher smoothness of continuously polynomial functions + +We prove that continuously polynomial functions are `C^∞`. In particular, this is the case +of continuous multilinear maps. +-/ + +open Filter Asymptotics + +open scoped ENNReal ContDiff + +universe u v + +variable {𝕜 : Type*} [NontriviallyNormedField 𝕜] +variable {E : Type u} [NormedAddCommGroup E] [NormedSpace 𝕜 E] +variable {F : Type v} [NormedAddCommGroup F] [NormedSpace 𝕜 F] + +section fderiv + +variable {p : FormalMultilinearSeries 𝕜 E F} {r : ℝ≥0∞} {n : ℕ} +variable {f : E → F} {x : E} {s : Set E} + +/-- A polynomial function is infinitely differentiable. -/ +theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : + ContDiffOn 𝕜 n f s := by + let t := { x | CPolynomialAt 𝕜 f x } + suffices ContDiffOn 𝕜 ω f t from (this.of_le le_top).mono h + rw [← contDiffOn_infty_iff_contDiffOn_omega] + have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx + have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f + exact contDiffOn_of_continuousOn_differentiableOn + (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr + fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) + +theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : + ContDiffAt 𝕜 n f x := + let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn + hf.contDiffOn.contDiffAt hs + +end fderiv + +namespace ContinuousMultilinearMap + +variable {ι : Type*} {E : ι → Type*} [∀ i, NormedAddCommGroup (E i)] [∀ i, NormedSpace 𝕜 (E i)] + [Fintype ι] (f : ContinuousMultilinearMap 𝕜 E F) {n : WithTop ℕ∞} {x : Π i, E i} + +open FormalMultilinearSeries + +lemma contDiffAt : ContDiffAt 𝕜 n f x := f.cpolynomialAt.contDiffAt + +lemma contDiff : ContDiff 𝕜 n f := contDiff_iff_contDiffAt.mpr (fun _ ↦ f.contDiffAt) + +end ContinuousMultilinearMap diff --git a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean index 9141ffab19f41..48cb70ce773fd 100644 --- a/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean +++ b/Mathlib/Analysis/Calculus/FDeriv/Analytic.lean @@ -7,7 +7,7 @@ import Mathlib.Analysis.Analytic.CPolynomial import Mathlib.Analysis.Analytic.Inverse import Mathlib.Analysis.Analytic.Within import Mathlib.Analysis.Calculus.Deriv.Basic -import Mathlib.Analysis.Calculus.ContDiff.Defs +import Mathlib.Analysis.Calculus.ContDiff.FTaylorSeries import Mathlib.Analysis.Calculus.FDeriv.Add import Mathlib.Analysis.Calculus.FDeriv.Prod import Mathlib.Analysis.Normed.Module.Completion @@ -63,7 +63,7 @@ differentiability at points in a neighborhood of `s`. Therefore, the theorem tha open Filter Asymptotics Set -open scoped ENNReal Topology ContDiff +open scoped ENNReal Topology universe u v @@ -192,6 +192,7 @@ theorem HasFPowerSeriesOnBall.fderiv_eq [CompleteSpace F] (h : HasFPowerSeriesOn fderiv 𝕜 f (x + y) = continuousMultilinearCurryFin1 𝕜 E F (p.changeOrigin y 1) := (h.hasFDerivAt hy).fderiv +/-- If a function has a power series on a ball, then so does its derivative. -/ protected theorem HasFPowerSeriesOnBall.fderiv [CompleteSpace F] (h : HasFPowerSeriesOnBall f p x r) : HasFPowerSeriesOnBall (fderiv 𝕜 f) p.derivSeries x r := by @@ -238,7 +239,7 @@ protected theorem AnalyticOnNhd.fderiv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 alias AnalyticOn.fderiv := AnalyticOnNhd.fderiv /-- If a function is analytic on a set `s`, so are its successive Fréchet derivative. See also -`AnalyticOnNhd.iteratedFDeruv_of_isOpen`, removing the completeness assumption but requiring the set +`AnalyticOnNhd.iteratedFDeriv_of_isOpen`, removing the completeness assumption but requiring the set to be open.-/ protected theorem AnalyticOnNhd.iteratedFDeriv [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) (n : ℕ) : AnalyticOnNhd 𝕜 (iteratedFDeriv 𝕜 n f) s := by @@ -269,44 +270,6 @@ lemma AnalyticOnNhd.hasFTaylorSeriesUpToOn [CompleteSpace F] · apply (DifferentiableAt.continuousAt (𝕜 := 𝕜) ?_).continuousWithinAt exact (h.iteratedFDeriv m x hx).differentiableAt -/-- An analytic function is infinitely differentiable. -/ -protected theorem AnalyticOnNhd.contDiffOn [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f s) - {n : WithTop ℕ∞} : ContDiffOn 𝕜 n f s := by - suffices ContDiffOn 𝕜 ω f s from this.of_le le_top - rw [← contDiffOn_infty_iff_contDiffOn_omega] - let t := { x | AnalyticAt 𝕜 f x } - suffices ContDiffOn 𝕜 ∞ f t from this.mono h - have H : AnalyticOnNhd 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_analyticAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (fun m _ ↦ (H.iteratedFDeriv m).differentiableOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - -/-- An analytic function on the whole space is infinitely differentiable there. -/ -theorem AnalyticOnNhd.contDiff [CompleteSpace F] (h : AnalyticOnNhd 𝕜 f univ) {n : WithTop ℕ∞} : - ContDiff 𝕜 n f := by - rw [← contDiffOn_univ] - exact h.contDiffOn - -theorem AnalyticAt.contDiffAt [CompleteSpace F] (h : AnalyticAt 𝕜 f x) {n : WithTop ℕ∞} : - ContDiffAt 𝕜 n f x := by - obtain ⟨s, hs, hf⟩ := h.exists_mem_nhds_analyticOnNhd - exact hf.contDiffOn.contDiffAt hs - -protected lemma AnalyticWithinAt.contDiffWithinAt [CompleteSpace F] {f : E → F} {s : Set E} {x : E} - (h : AnalyticWithinAt 𝕜 f s x) {n : ℕ∞} : ContDiffWithinAt 𝕜 n f s x := by - rcases h.exists_analyticAt with ⟨g, fx, fg, hg⟩ - exact hg.contDiffAt.contDiffWithinAt.congr (fg.mono (subset_insert _ _)) fx - -protected lemma AnalyticOn.contDiffOn [CompleteSpace F] {f : E → F} {s : Set E} - (h : AnalyticOn 𝕜 f s) {n : ℕ∞} : ContDiffOn 𝕜 n f s := - fun x m ↦ (h x m).contDiffWithinAt - -@[deprecated (since := "2024-09-26")] -alias AnalyticWithinOn.contDiffOn := AnalyticOn.contDiffOn - lemma AnalyticWithinAt.exists_hasFTaylorSeriesUpToOn [CompleteSpace F] (n : WithTop ℕ∞) (h : AnalyticWithinAt 𝕜 f s x) : ∃ u ∈ 𝓝[insert x s] x, ∃ (p : E → FormalMultilinearSeries 𝕜 E F), @@ -384,7 +347,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 : WithTop ℕ∞} +protected 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 ↦ ?_⟩ @@ -546,26 +509,6 @@ theorem CPolynomialOn.iteratedFDeriv (h : CPolynomialOn 𝕜 f s) (n : ℕ) : case g => exact ↑(continuousMultilinearCurryLeftEquiv 𝕜 (fun _ : Fin (n + 1) ↦ E) F).symm simp -/-- A polynomial function is infinitely differentiable. -/ -theorem CPolynomialOn.contDiffOn (h : CPolynomialOn 𝕜 f s) {n : WithTop ℕ∞} : - ContDiffOn 𝕜 n f s := by - suffices ContDiffOn 𝕜 ω f s from this.of_le le_top - let t := { x | CPolynomialAt 𝕜 f x } - suffices ContDiffOn 𝕜 ω f t from this.mono h - rw [← contDiffOn_infty_iff_contDiffOn_omega] - have H : CPolynomialOn 𝕜 f t := fun _x hx ↦ hx - have t_open : IsOpen t := isOpen_cPolynomialAt 𝕜 f - exact contDiffOn_of_continuousOn_differentiableOn - (fun m _ ↦ (H.iteratedFDeriv m).continuousOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - (fun m _ ↦ (H.iteratedFDeriv m).analyticOnNhd.differentiableOn.congr - fun _ hx ↦ iteratedFDerivWithin_of_isOpen _ t_open hx) - -theorem CPolynomialAt.contDiffAt (h : CPolynomialAt 𝕜 f x) {n : WithTop ℕ∞} : - ContDiffAt 𝕜 n f x := - let ⟨_, hs, hf⟩ := h.exists_mem_nhds_cPolynomialOn - hf.contDiffOn.contDiffAt hs - end fderiv section deriv @@ -750,10 +693,6 @@ lemma cPolynomialAt : CPolynomialAt 𝕜 f x := lemma cPolyomialOn : CPolynomialOn 𝕜 f ⊤ := fun x _ ↦ f.cPolynomialAt x -lemma contDiffAt : ContDiffAt 𝕜 n f x := (f.cPolynomialAt x).contDiffAt - -lemma contDiff : ContDiff 𝕜 n f := contDiff_iff_contDiffAt.mpr f.contDiffAt - end ContinuousMultilinearMap namespace FormalMultilinearSeries @@ -797,8 +736,8 @@ private theorem factorial_smul' {n : ℕ} : ∀ {F : Type max u v} [NormedAddCom n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by induction n with | zero => _ | succ n ih => _ <;> intro F _ _ _ p f h · rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag] - · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply, - ih h.fderiv, iteratedFDeriv_succ_apply_right] + · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, + ← ContinuousLinearMap.smul_apply, ih h.fderiv, iteratedFDeriv_succ_apply_right] rfl variable [CompleteSpace F] @@ -808,8 +747,8 @@ theorem factorial_smul (n : ℕ) : n ! • p n (fun _ ↦ y) = iteratedFDeriv 𝕜 n f x (fun _ ↦ y) := by cases n · rw [factorial_zero, one_smul, h.iteratedFDeriv_zero_apply_diag] - · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, ← smul_apply, - factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right] + · rw [factorial_succ, mul_comm, mul_smul, ← derivSeries_apply_diag, + ← ContinuousLinearMap.smul_apply, factorial_smul' _ h.fderiv, iteratedFDeriv_succ_apply_right] rfl theorem hasSum_iteratedFDeriv [CharZero 𝕜] {y : E} (hy : y ∈ EMetric.ball 0 r) : diff --git a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean index 0151b54688137..dcf554692ab17 100644 --- a/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean +++ b/Mathlib/Analysis/Fourier/FourierTransformDeriv.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Alex Kontorovich, David Loeffler, Heather Macbeth, Sébastien Gouëzel -/ import Mathlib.Analysis.Calculus.ParametricIntegral +import Mathlib.Analysis.Calculus.ContDiff.CPolynomial import Mathlib.Analysis.Fourier.AddCircle import Mathlib.Analysis.Fourier.FourierTransform import Mathlib.Analysis.Calculus.FDeriv.Analytic @@ -369,7 +370,7 @@ lemma norm_iteratedFDeriv_fourierPowSMulRight have I₂ m : ‖iteratedFDeriv ℝ m (T ∘ (ContinuousLinearMap.pi (fun (_ : Fin n) ↦ L))) v‖ ≤ (n.descFactorial m * 1 * (‖L‖ * ‖v‖) ^ (n - m)) * ‖L‖ ^ m := by rw [ContinuousLinearMap.iteratedFDeriv_comp_right _ (ContinuousMultilinearMap.contDiff _) - _ le_top] + _ (mod_cast le_top)] apply (norm_compContinuousLinearMap_le _ _).trans simp only [Finset.prod_const, Finset.card_fin] gcongr diff --git a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean index 846339ecb2f00..7a9ffe7ff6f72 100644 --- a/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean +++ b/Mathlib/Analysis/SpecialFunctions/ExpDeriv.lean @@ -4,7 +4,9 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ import Mathlib.Analysis.Complex.RealDeriv +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Analysis.Calculus.ContDiff.RCLike +import Mathlib.Analysis.Calculus.FDeriv.Analytic import Mathlib.Analysis.Calculus.IteratedDeriv.Lemmas import Mathlib.Analysis.SpecialFunctions.Exponential diff --git a/Mathlib/Geometry/Manifold/AnalyticManifold.lean b/Mathlib/Geometry/Manifold/AnalyticManifold.lean index 9a1ab30e1daa9..ea26292a58a0d 100644 --- a/Mathlib/Geometry/Manifold/AnalyticManifold.lean +++ b/Mathlib/Geometry/Manifold/AnalyticManifold.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Michael Lee, Geoffrey Irving -/ import Mathlib.Analysis.Analytic.Constructions -import Mathlib.Analysis.Calculus.FDeriv.Analytic +import Mathlib.Analysis.Calculus.ContDiff.Analytic import Mathlib.Geometry.Manifold.SmoothManifoldWithCorners /-! From e10334bd8b2298ee04df07fbd1896ffb60e143b6 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 09:04:46 +0000 Subject: [PATCH 213/250] =?UTF-8?q?chore:=20rename=20`div=5Fdiv=5Fcancel'`?= =?UTF-8?q?=20to=20`div=5Fdiv=5Fcancel=E2=82=80`=20(#19459)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit and make it simp --- Mathlib/Algebra/GroupWithZero/Units/Basic.lean | 4 +++- Mathlib/Algebra/Order/Field/Basic.lean | 2 +- Mathlib/Analysis/Complex/Liouville.lean | 2 +- Mathlib/Analysis/Complex/PhragmenLindelof.lean | 2 +- Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean | 2 +- 5 files changed, 7 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index be6aebca0a776..6dcca5a5b9844 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -451,7 +451,9 @@ lemma div_eq_div_iff_div_eq_div' (hb : b ≠ 0) (hc : c ≠ 0) : a / b = c / d conv_rhs => rw [← mul_left_inj' hc, div_mul_cancel₀ _ hc] rw [mul_comm _ c, div_mul_eq_mul_div, mul_div_assoc] -lemma div_div_cancel' (ha : a ≠ 0) : a / (a / b) = b := ha.isUnit.div_div_cancel +@[simp] lemma div_div_cancel₀ (ha : a ≠ 0) : a / (a / b) = b := ha.isUnit.div_div_cancel + +@[deprecated (since := "2024-11-25")] alias div_div_cancel' := div_div_cancel₀ lemma div_div_cancel_left' (ha : a ≠ 0) : a / b / a = b⁻¹ := ha.isUnit.div_div_cancel_left diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index 5832b57160d05..a8faa3f8b1daf 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -344,7 +344,7 @@ theorem div_mul_le_div_mul_of_div_le_div (h : a / b ≤ c / d) (he : 0 ≤ e) : theorem exists_pos_mul_lt {a : α} (h : 0 < a) (b : α) : ∃ c : α, 0 < c ∧ b * c < a := by have : 0 < a / max (b + 1) 1 := div_pos h (lt_max_iff.2 (Or.inr zero_lt_one)) refine ⟨a / max (b + 1) 1, this, ?_⟩ - rw [← lt_div_iff₀ this, div_div_cancel' h.ne'] + rw [← lt_div_iff₀ this, div_div_cancel₀ h.ne'] exact lt_max_iff.2 (Or.inl <| lt_add_one _) theorem exists_pos_lt_mul {a : α} (h : 0 < a) (b : α) : ∃ c : α, 0 < c ∧ b < c * a := diff --git a/Mathlib/Analysis/Complex/Liouville.lean b/Mathlib/Analysis/Complex/Liouville.lean index 2d327b794e416..aee36657dab9e 100644 --- a/Mathlib/Analysis/Complex/Liouville.lean +++ b/Mathlib/Analysis/Complex/Liouville.lean @@ -93,7 +93,7 @@ theorem liouville_theorem_aux {f : ℂ → F} (hf : Differentiable ℂ f) (hb : calc ‖deriv f c‖ ≤ C / (C / ε) := norm_deriv_le_of_forall_mem_sphere_norm_le (div_pos C₀ ε₀) hf.diffContOnCl fun z _ => hC z - _ = ε := div_div_cancel' C₀.lt.ne' + _ = ε := div_div_cancel₀ C₀.lt.ne' end Complex diff --git a/Mathlib/Analysis/Complex/PhragmenLindelof.lean b/Mathlib/Analysis/Complex/PhragmenLindelof.lean index a457d182cf842..ca347d8adbcac 100644 --- a/Mathlib/Analysis/Complex/PhragmenLindelof.lean +++ b/Mathlib/Analysis/Complex/PhragmenLindelof.lean @@ -362,7 +362,7 @@ nonrec theorem quadrant_I (hd : DiffContOnCl ℂ f (Ioi 0 ×ℂ Ioi 0)) -- Porting note: failed to clear hζ ζ · -- The estimate `hB` on `f` implies the required estimate on -- `f ∘ exp` with the same `c` and `B' = max B 0`. - rw [sub_zero, div_div_cancel' Real.pi_pos.ne'] + rw [sub_zero, div_div_cancel₀ Real.pi_pos.ne'] rcases hB with ⟨c, hc, B, hO⟩ refine ⟨c, hc, max B 0, ?_⟩ rw [← comap_comap, comap_abs_atTop, comap_sup, inf_sup_right] diff --git a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean index 493dc137e80c5..0dade7c5124a1 100644 --- a/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean +++ b/Mathlib/NumberTheory/LSeries/AbstractFuncEq.lean @@ -103,7 +103,7 @@ section symmetry /-- Reformulated functional equation with `f` and `g` interchanged. -/ lemma WeakFEPair.h_feq' (P : WeakFEPair E) (x : ℝ) (hx : 0 < x) : P.g (1 / x) = (P.ε⁻¹ * ↑(x ^ P.k)) • P.f x := by - rw [(div_div_cancel' (one_ne_zero' ℝ) ▸ P.h_feq (1 / x) (one_div_pos.mpr hx):), ← mul_smul] + rw [(div_div_cancel₀ (one_ne_zero' ℝ) ▸ P.h_feq (1 / x) (one_div_pos.mpr hx):), ← mul_smul] convert (one_smul ℂ (P.g (1 / x))).symm using 2 rw [one_div, inv_rpow hx.le, ofReal_inv] field_simp [P.hε, (rpow_pos_of_pos hx _).ne'] From 092a7d66c70504960126d34e8cc455dd042ddd91 Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Nov 2024 09:56:40 +0000 Subject: [PATCH 214/250] feat: coproducts in the category of Ind-objects (#19154) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Shows that - the functor `C ⥤ Ind C` preserves finite colimits, - if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α`, - if `C` has finite coproducts, then `Ind C` has small coproducts. Co-authored-by: Markus Himmel --- .../Limits/Indization/Category.lean | 83 ++++++++++++++++++- 1 file changed, 80 insertions(+), 3 deletions(-) diff --git a/Mathlib/CategoryTheory/Limits/Indization/Category.lean b/Mathlib/CategoryTheory/Limits/Indization/Category.lean index c0778f14edc72..c927eac12564d 100644 --- a/Mathlib/CategoryTheory/Limits/Indization/Category.lean +++ b/Mathlib/CategoryTheory/Limits/Indization/Category.lean @@ -3,9 +3,11 @@ Copyright (c) 2024 Markus Himmel. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Markus Himmel -/ +import Mathlib.CategoryTheory.Functor.Flat +import Mathlib.CategoryTheory.Limits.Constructions.Filtered +import Mathlib.CategoryTheory.Limits.FullSubcategory import Mathlib.CategoryTheory.Limits.Indization.LocallySmall import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits -import Mathlib.CategoryTheory.Limits.FullSubcategory /-! # The category of Ind-objects @@ -14,8 +16,22 @@ We define the `v`-category of Ind-objects of a category `C`, called `Ind C`, as `Ind.yoneda : C ⥤ Ind C` and `Ind.inclusion C : Ind C ⥤ Cᵒᵖ ⥤ Type v`. This file will mainly collect results about ind-objects (stated in terms of `IsIndObject`) and -reinterpret them in terms of `Ind C`. For now, we show that `Ind C` has filtered colimits and that -`Ind.inclusion C` creates them. Many other limit-colimit properties will follow. +reinterpret them in terms of `Ind C`. + +Adopting the theorem numbering of [Kashiwara2006], we show that +* `Ind C` has filtered colimits (6.1.8), +* `Ind C ⥤ Cᵒᵖ ⥤ Type v` creates filtered colimits (6.1.8), +* `C ⥤ Ind C` preserves finite colimits (6.1.6), +* if `C` has coproducts indexed by a finite type `α`, then `Ind C` has coproducts indexed by `α` + (6.1.18(ii)), +* if `C` has finite coproducts, then `Ind C` has small coproducts (6.1.18(ii)). + +More limit-colimit properties will follow. + +Note that: +* the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` does not preserve any kind of colimit in general except for + filtered colimits and +* the functor `C ⥤ Ind C` preserves finite colimits, but not infinite colimits in general. ## References * [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Chapter 6 @@ -54,6 +70,10 @@ instance : (Ind.inclusion C).Full := instance : (Ind.inclusion C).Faithful := inferInstanceAs <| ((Ind.equivalence C).functor ⋙ fullSubcategoryInclusion _).Faithful +/-- The functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` is fully faithful. -/ +protected noncomputable def Ind.inclusion.fullyFaithful : (Ind.inclusion C).FullyFaithful := + .ofFullyFaithful _ + /-- The inclusion of `C` into `Ind C` induced by the Yoneda embedding. -/ protected noncomputable def Ind.yoneda : C ⥤ Ind C := FullSubcategory.lift _ CategoryTheory.yoneda isIndObject_yoneda ⋙ (Ind.equivalence C).inverse @@ -66,6 +86,10 @@ instance : (Ind.yoneda (C := C)).Faithful := inferInstanceAs <| Functor.Faithful <| FullSubcategory.lift _ CategoryTheory.yoneda isIndObject_yoneda ⋙ (Ind.equivalence C).inverse +/-- The functor `C ⥤ Ind C` is fully faithful. -/ +protected noncomputable def Ind.yoneda.fullyFaithful : (Ind.yoneda (C := C)).FullyFaithful := + .ofFullyFaithful _ + /-- The composition `C ⥤ Ind C ⥤ (Cᵒᵖ ⥤ Type v)` is just the Yoneda embedding. -/ noncomputable def Ind.yonedaCompInclusion : Ind.yoneda ⋙ Ind.inclusion C ≅ CategoryTheory.yoneda := isoWhiskerLeft (FullSubcategory.lift _ _ _) @@ -83,4 +107,57 @@ instance : HasFilteredColimits (Ind C) where HasColimitsOfShape _ _ _ := hasColimitsOfShape_of_hasColimitsOfShape_createsColimitsOfShape (Ind.inclusion C) +theorem Ind.isIndObject_inclusion_obj (X : Ind C) : IsIndObject ((Ind.inclusion C).obj X) := + X.2 + +/-- Pick a presentation of an ind-object `X` using choice. -/ +noncomputable def Ind.presentation (X : Ind C) : IndObjectPresentation ((Ind.inclusion C).obj X) := + X.isIndObject_inclusion_obj.presentation + +/-- An ind-object `X` is the colimit (in `Ind C`!) of the filtered diagram presenting it. -/ +noncomputable def Ind.colimitPresentationCompYoneda (X : Ind C) : + colimit (X.presentation.F ⋙ Ind.yoneda) ≅ X := + Ind.inclusion.fullyFaithful.isoEquiv.symm <| calc + (Ind.inclusion C).obj (colimit (X.presentation.F ⋙ Ind.yoneda)) + ≅ colimit (X.presentation.F ⋙ Ind.yoneda ⋙ Ind.inclusion C) := preservesColimitIso _ _ + _ ≅ colimit (X.presentation.F ⋙ yoneda) := + HasColimit.isoOfNatIso (isoWhiskerLeft X.presentation.F Ind.yonedaCompInclusion) + _ ≅ (Ind.inclusion C).obj X := + IsColimit.coconePointUniqueUpToIso (colimit.isColimit _) X.presentation.isColimit + +instance : RepresentablyCoflat (Ind.yoneda (C := C)) := by + refine ⟨fun X => ?_⟩ + suffices IsFiltered (CostructuredArrow yoneda ((Ind.inclusion C).obj X)) from + IsFiltered.of_equivalence + ((CostructuredArrow.post Ind.yoneda (Ind.inclusion C) X).asEquivalence.trans + (CostructuredArrow.mapNatIso Ind.yonedaCompInclusion)).symm + exact ((isIndObject_iff _).1 (Ind.isIndObject_inclusion_obj X)).1 + +noncomputable instance : PreservesFiniteColimits (Ind.yoneda (C := C)) := + preservesFiniteColimits_of_coflat _ + +instance {α : Type v} [Finite α] [HasColimitsOfShape (Discrete α) C] : + HasColimitsOfShape (Discrete α) (Ind C) := by + refine ⟨fun F => ?_⟩ + let I : α → Type v := fun s => (F.obj ⟨s⟩).presentation.I + let G : ∀ s, I s ⥤ C := fun s => (F.obj ⟨s⟩).presentation.F + let iso : Discrete.functor (fun s => Pi.eval I s ⋙ G s) ⋙ + (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim ≅ F := by + refine Discrete.natIso (fun s => ?_) + refine (Functor.Final.colimitIso (Pi.eval I s.as) (G s.as ⋙ Ind.yoneda)) ≪≫ ?_ + exact Ind.colimitPresentationCompYoneda _ + -- The actual proof happens during typeclass resolution in the following line, which deduces + -- ``` + -- HasColimit Discrete.functor (fun s => Pi.eval I s ⋙ G s) ⋙ + -- (whiskeringRight _ _ _).obj Ind.yoneda ⋙ colim + -- ``` + -- from the fact that finite limits commute with filtered colimits and from the fact that + -- `Ind.yoneda` preserves finite colimits. + apply hasColimitOfIso iso.symm + +instance [HasFiniteCoproducts C] : HasCoproducts.{v} (Ind C) := + have : HasFiniteCoproducts (Ind C) := + ⟨fun _ => hasColimitsOfShape_of_equivalence (Discrete.equivalence Equiv.ulift)⟩ + hasCoproducts_of_finite_and_filtered + end CategoryTheory From 134190977464360ecc5de1cd85ffb25b1172db7d Mon Sep 17 00:00:00 2001 From: grhkm21 Date: Wed, 27 Nov 2024 09:56:41 +0000 Subject: [PATCH 215/250] feat: prove `zmodRepr` is unique (#19391) We prove `PadicInt.exists_unique_mem_range` which is a strengthening of `exists_mem_range`, proving uniqueness of the natural number. I need this to simplify `toZMod(1+ap)` to `1`. --- Mathlib/NumberTheory/Padics/RingHoms.lean | 14 ++++++++++++-- 1 file changed, 12 insertions(+), 2 deletions(-) diff --git a/Mathlib/NumberTheory/Padics/RingHoms.lean b/Mathlib/NumberTheory/Padics/RingHoms.lean index b80610a74c52a..2fed1c971dbfe 100644 --- a/Mathlib/NumberTheory/Padics/RingHoms.lean +++ b/Mathlib/NumberTheory/Padics/RingHoms.lean @@ -178,14 +178,24 @@ theorem exists_mem_range : ∃ n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] apply max_lt hr simpa using hn +theorem exists_unique_mem_range : ∃! n : ℕ, n < p ∧ x - n ∈ maximalIdeal ℤ_[p] := by + obtain ⟨n, hn₁, hn₂⟩ := exists_mem_range x + use n, ⟨hn₁, hn₂⟩, fun m ⟨hm₁, hm₂⟩ ↦ ?_ + have := (zmod_congr_of_sub_mem_max_ideal x n m hn₂ hm₂).symm + rwa [ZMod.natCast_eq_natCast_iff, ModEq, mod_eq_of_lt hn₁, mod_eq_of_lt hm₁] at this + /-- `zmod_repr x` is the unique natural number smaller than `p` satisfying `‖(x - zmod_repr x : ℤ_[p])‖ < 1`. -/ def zmodRepr : ℕ := - Classical.choose (exists_mem_range x) + Classical.choose (exists_unique_mem_range x).exists theorem zmodRepr_spec : zmodRepr x < p ∧ x - zmodRepr x ∈ maximalIdeal ℤ_[p] := - Classical.choose_spec (exists_mem_range x) + Classical.choose_spec (exists_unique_mem_range x).exists + +theorem zmodRepr_unique (y : ℕ) (hy₁ : y < p) (hy₂ : x - y ∈ maximalIdeal ℤ_[p]) : y = zmodRepr x := + have h := (Classical.choose_spec (exists_unique_mem_range x)).right + (h y ⟨hy₁, hy₂⟩).trans (h (zmodRepr x) (zmodRepr_spec x)).symm theorem zmodRepr_lt_p : zmodRepr x < p := (zmodRepr_spec _).1 From 8674145a5b085f53aafb08bca0fd7eae8e9a1d9d Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 09:56:42 +0000 Subject: [PATCH 216/250] chore(NumberTheory/Fermat): fix theorem names (#19469) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit From LeanCamCombi Co-authored-by: Yves Jäckle --- Mathlib/NumberTheory/Fermat.lean | 49 ++++++++++++++++++++------------ 1 file changed, 31 insertions(+), 18 deletions(-) diff --git a/Mathlib/NumberTheory/Fermat.lean b/Mathlib/NumberTheory/Fermat.lean index bdb8dd1117803..c2aca1d62719d 100644 --- a/Mathlib/NumberTheory/Fermat.lean +++ b/Mathlib/NumberTheory/Fermat.lean @@ -20,6 +20,8 @@ for all natural numbers `n`. form `k * 2 ^ (n + 2) + 1`. -/ +open Function + namespace Nat open Finset Nat ZMod @@ -32,20 +34,25 @@ def fermatNumber (n : ℕ) : ℕ := 2 ^ (2 ^ n) + 1 @[simp] theorem fermatNumber_one : fermatNumber 1 = 5 := rfl @[simp] theorem fermatNumber_two : fermatNumber 2 = 17 := rfl -theorem strictMono_fermatNumber : StrictMono fermatNumber := by +theorem fermatNumber_strictMono : StrictMono fermatNumber := by intro m n simp only [fermatNumber, add_lt_add_iff_right, Nat.pow_lt_pow_iff_right (one_lt_two : 1 < 2), imp_self] -theorem two_lt_fermatNumber (n : ℕ) : 2 < fermatNumber n := by - cases n - · simp - · exact lt_of_succ_lt <| strictMono_fermatNumber <| zero_lt_succ _ +@[deprecated (since := "2024-11-25")] alias strictMono_fermatNumber := fermatNumber_strictMono + +lemma fermatNumber_mono : Monotone fermatNumber := fermatNumber_strictMono.monotone +lemma fermatNumber_injective : Injective fermatNumber := fermatNumber_strictMono.injective + +lemma three_le_fermatNumber (n : ℕ) : 3 ≤ fermatNumber n := fermatNumber_mono n.zero_le +lemma two_lt_fermatNumber (n : ℕ) : 2 < fermatNumber n := three_le_fermatNumber _ + +lemma fermatNumber_ne_one (n : ℕ) : fermatNumber n ≠ 1 := by have := three_le_fermatNumber n; omega theorem odd_fermatNumber (n : ℕ) : Odd (fermatNumber n) := (even_pow.mpr ⟨even_two, (pow_pos two_pos n).ne'⟩).add_one -theorem fermatNumber_product (n : ℕ) : ∏ k in range n, fermatNumber k = fermatNumber n - 2 := by +theorem prod_fermatNumber (n : ℕ) : ∏ k ∈ range n, fermatNumber k = fermatNumber n - 2 := by induction' n with n hn · rfl rw [prod_range_succ, hn, fermatNumber, fermatNumber, mul_comm, @@ -53,9 +60,11 @@ theorem fermatNumber_product (n : ℕ) : ∏ k in range n, fermatNumber k = ferm ring_nf omega +@[deprecated (since := "2024-11-25")] alias fermatNumber_product := prod_fermatNumber + theorem fermatNumber_eq_prod_add_two (n : ℕ) : - fermatNumber n = (∏ k in range n, fermatNumber k) + 2 := by - rw [fermatNumber_product, Nat.sub_add_cancel] + fermatNumber n = ∏ k ∈ range n, fermatNumber k + 2 := by + rw [prod_fermatNumber, Nat.sub_add_cancel] exact le_of_lt <| two_lt_fermatNumber _ theorem fermatNumber_succ (n : ℕ) : fermatNumber (n + 1) = (fermatNumber n - 1) ^ 2 + 1 := by @@ -93,16 +102,20 @@ open Finset From a letter to Euler, see page 37 in [juskevic2022]. -/ -theorem coprime_fermatNumber_fermatNumber {k n : ℕ} (h : k ≠ n) : - Coprime (fermatNumber n) (fermatNumber k) := by - wlog hkn : k < n - · simpa only [coprime_comm] using this h.symm (by omega) - let m := (fermatNumber n).gcd (fermatNumber k) - have h_n : m ∣ fermatNumber n := (fermatNumber n).gcd_dvd_left (fermatNumber k) - have h_m : m ∣ 2 := (Nat.dvd_add_right <| (gcd_dvd_right _ _).trans <| dvd_prod_of_mem _ - <| mem_range.mpr hkn).mp <| fermatNumber_eq_prod_add_two _ ▸ h_n - refine ((dvd_prime prime_two).mp h_m).elim id (fun h_two ↦ ?_) - exact ((odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n)).elim +theorem coprime_fermatNumber_fermatNumber {m n : ℕ} (hmn : m ≠ n) : + Coprime (fermatNumber m) (fermatNumber n) := by + wlog hmn' : m < n + · simpa only [coprime_comm] using this hmn.symm (by omega) + let d := (fermatNumber m).gcd (fermatNumber n) + have h_n : d ∣ fermatNumber n := gcd_dvd_right .. + have h_m : d ∣ 2 := (Nat.dvd_add_right <| (gcd_dvd_left _ _).trans <| dvd_prod_of_mem _ + <| mem_range.mpr hmn').mp <| fermatNumber_eq_prod_add_two _ ▸ h_n + refine ((dvd_prime prime_two).mp h_m).resolve_right fun h_two ↦ ?_ + exact (odd_fermatNumber _).not_two_dvd_nat (h_two ▸ h_n) + +lemma pairwise_coprime_fermatNumber : + Pairwise fun m n ↦ Coprime (fermatNumber m) (fermatNumber n) := + fun _m _n ↦ coprime_fermatNumber_fermatNumber open ZMod From 00200113ae7e6500fe14d089684b57daef691d0e Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 09:56:44 +0000 Subject: [PATCH 217/250] chore(GroupTheory/FreeGroup): split off `Reduce.lean` (#19501) --- Mathlib.lean | 1 + Mathlib/GroupTheory/FreeGroup/Basic.lean | 326 +------------------- Mathlib/GroupTheory/FreeGroup/Reduce.lean | 351 ++++++++++++++++++++++ Mathlib/SetTheory/Cardinal/Free.lean | 1 + 4 files changed, 354 insertions(+), 325 deletions(-) create mode 100644 Mathlib/GroupTheory/FreeGroup/Reduce.lean diff --git a/Mathlib.lean b/Mathlib.lean index 21358bdd55951..ff6ae6de6f371 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3115,6 +3115,7 @@ import Mathlib.GroupTheory.FreeAbelianGroupFinsupp import Mathlib.GroupTheory.FreeGroup.Basic import Mathlib.GroupTheory.FreeGroup.IsFreeGroup import Mathlib.GroupTheory.FreeGroup.NielsenSchreier +import Mathlib.GroupTheory.FreeGroup.Reduce import Mathlib.GroupTheory.GroupAction.Basic import Mathlib.GroupTheory.GroupAction.Blocks import Mathlib.GroupTheory.GroupAction.CardCommute diff --git a/Mathlib/GroupTheory/FreeGroup/Basic.lean b/Mathlib/GroupTheory/FreeGroup/Basic.lean index 2c47200b1ac4d..a2174bda20dfc 100644 --- a/Mathlib/GroupTheory/FreeGroup/Basic.lean +++ b/Mathlib/GroupTheory/FreeGroup/Basic.lean @@ -3,9 +3,8 @@ 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.BigOperators.Group.List import Mathlib.Algebra.Group.Subgroup.Ker -import Mathlib.Data.Fintype.Basic -import Mathlib.Data.List.Sublists /-! # Free groups @@ -895,327 +894,4 @@ instance : LawfulMonad FreeGroup.{u} := LawfulMonad.mk' end Category -section Reduce - -variable [DecidableEq α] - -/-- The maximal reduction of a word. It is computable -iff `α` has decidable equality. -/ -@[to_additive "The maximal reduction of a word. It is computable iff `α` has decidable equality."] -def reduce : (L : List (α × Bool)) -> List (α × Bool) := - List.rec [] fun hd1 _tl1 ih => - List.casesOn ih [hd1] fun hd2 tl2 => - if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 - -@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl -@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl - -@[to_additive (attr := simp)] -theorem reduce.cons (x) : - reduce (x :: L) = - List.casesOn (reduce L) [x] fun hd tl => - if x.1 = hd.1 ∧ x.2 = not hd.2 then tl else x :: hd :: tl := - rfl - -@[to_additive (attr := simp)] -theorem reduce_replicate (n : ℕ) (x : α × Bool) : - reduce (.replicate n x) = .replicate n x := by - induction n with - | zero => simp [reduce] - | succ n ih => - rw [List.replicate_succ, reduce.cons, ih] - cases n with - | zero => simp - | succ n => simp [List.replicate_succ] - -/-- The first theorem that characterises the function `reduce`: a word reduces to its maximal - reduction. -/ -@[to_additive "The first theorem that characterises the function `reduce`: a word reduces to its - maximal reduction."] -theorem reduce.red : Red L (reduce L) := by - induction L with - | nil => constructor - | cons hd1 tl1 ih => - dsimp - revert ih - generalize htl : reduce tl1 = TL - intro ih - cases TL with - | nil => exact Red.cons_cons ih - | cons hd2 tl2 => - dsimp only - split_ifs with h - · cases hd1 - cases hd2 - cases h - dsimp at * - subst_vars - apply Red.trans (Red.cons_cons ih) - exact Red.Step.cons_not_rev.to_red - · exact Red.cons_cons ih - -@[to_additive] -theorem reduce.not {p : Prop} : - ∀ {L₁ L₂ L₃ : List (α × Bool)} {x b}, reduce L₁ = L₂ ++ (x, b) :: (x, !b) :: L₃ → p - | [], L2, L3, _, _ => fun h => by cases L2 <;> injections - | (x, b) :: L1, L2, L3, x', b' => by - dsimp - cases r : reduce L1 with - | nil => - dsimp; intro h - exfalso - have := congr_arg List.length h - simp? [List.length] at this says - simp only [List.length, zero_add, List.length_append] at this - rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this - omega - | cons hd tail => - cases' hd with y c - dsimp only - split_ifs with h <;> intro H - · rw [H] at r - exact @reduce.not _ L1 ((y, c) :: L2) L3 x' b' r - rcases L2 with (_ | ⟨a, L2⟩) - · injections; subst_vars - simp at h - · refine @reduce.not _ L1 L2 L3 x' b' ?_ - injection H with _ H - rw [r, H]; rfl - -/-- The second theorem that characterises the function `reduce`: the maximal reduction of a word -only reduces to itself. -/ -@[to_additive "The second theorem that characterises the function `reduce`: the maximal reduction of - a word only reduces to itself."] -theorem reduce.min (H : Red (reduce L₁) L₂) : reduce L₁ = L₂ := by - induction' H with L1 L' L2 H1 H2 ih - · rfl - · cases' H1 with L4 L5 x b - exact reduce.not H2 - -/-- `reduce` is idempotent, i.e. the maximal reduction of the maximal reduction of a word is the - maximal reduction of the word. -/ -@[to_additive (attr := simp) "`reduce` is idempotent, i.e. the maximal reduction of the maximal - reduction of a word is the maximal reduction of the word."] -theorem reduce.idem : reduce (reduce L) = reduce L := - Eq.symm <| reduce.min reduce.red - -@[to_additive] -theorem reduce.Step.eq (H : Red.Step L₁ L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (reduce.red.head H) - (reduce.min HR13).trans (reduce.min HR23).symm - -/-- If a word reduces to another word, then they have a common maximal reduction. -/ -@[to_additive "If a word reduces to another word, then they have a common maximal reduction."] -theorem reduce.eq_of_red (H : Red L₁ L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (Red.trans H reduce.red) - (reduce.min HR13).trans (reduce.min HR23).symm - -alias red.reduce_eq := reduce.eq_of_red - -alias freeAddGroup.red.reduce_eq := FreeAddGroup.reduce.eq_of_red - -@[to_additive] -theorem Red.reduce_right (h : Red L₁ L₂) : Red L₁ (reduce L₂) := - reduce.eq_of_red h ▸ reduce.red - -@[to_additive] -theorem Red.reduce_left (h : Red L₁ L₂) : Red L₂ (reduce L₁) := - (reduce.eq_of_red h).symm ▸ reduce.red - -/-- If two words correspond to the same element in the free group, then they -have a common maximal reduction. This is the proof that the function that sends -an element of the free group to its maximal reduction is well-defined. -/ -@[to_additive "If two words correspond to the same element in the additive free group, then they - have a common maximal reduction. This is the proof that the function that sends an element of the - free group to its maximal reduction is well-defined."] -theorem reduce.sound (H : mk L₁ = mk L₂) : reduce L₁ = reduce L₂ := - let ⟨_L₃, H13, H23⟩ := Red.exact.1 H - (reduce.eq_of_red H13).trans (reduce.eq_of_red H23).symm - -/-- If two words have a common maximal reduction, then they correspond to the same element in the - free group. -/ -@[to_additive "If two words have a common maximal reduction, then they correspond to the same - element in the additive free group."] -theorem reduce.exact (H : reduce L₁ = reduce L₂) : mk L₁ = mk L₂ := - Red.exact.2 ⟨reduce L₂, H ▸ reduce.red, reduce.red⟩ - -/-- A word and its maximal reduction correspond to the same element of the free group. -/ -@[to_additive "A word and its maximal reduction correspond to the same element of the additive free - group."] -theorem reduce.self : mk (reduce L) = mk L := - reduce.exact reduce.idem - -/-- If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal reduction - of `w₁`. -/ -@[to_additive "If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal - reduction of `w₁`."] -theorem reduce.rev (H : Red L₁ L₂) : Red L₂ (reduce L₁) := - (reduce.eq_of_red H).symm ▸ reduce.red - -/-- The function that sends an element of the free group to its maximal reduction. -/ -@[to_additive "The function that sends an element of the additive free group to its maximal - reduction."] -def toWord : FreeGroup α → List (α × Bool) := - Quot.lift reduce fun _L₁ _L₂ H => reduce.Step.eq H - -@[to_additive] -theorem mk_toWord : ∀ {x : FreeGroup α}, mk (toWord x) = x := by rintro ⟨L⟩; exact reduce.self - -@[to_additive] -theorem toWord_injective : Function.Injective (toWord : FreeGroup α → List (α × Bool)) := by - rintro ⟨L₁⟩ ⟨L₂⟩; exact reduce.exact - -@[to_additive (attr := simp)] -theorem toWord_inj {x y : FreeGroup α} : toWord x = toWord y ↔ x = y := - toWord_injective.eq_iff - -@[to_additive (attr := simp)] -theorem toWord_mk : (mk L₁).toWord = reduce L₁ := - rfl - -@[to_additive (attr := simp)] -theorem toWord_of (a : α) : (of a).toWord = [(a, true)] := - rfl - -@[to_additive (attr := simp)] -theorem reduce_toWord : ∀ x : FreeGroup α, reduce (toWord x) = toWord x := by - rintro ⟨L⟩ - exact reduce.idem - -@[to_additive (attr := simp)] -theorem toWord_one : (1 : FreeGroup α).toWord = [] := - rfl - -@[to_additive (attr := simp)] -theorem toWord_of_pow (a : α) (n : ℕ) : (of a ^ n).toWord = List.replicate n (a, true) := by - rw [of, pow_mk, List.flatten_replicate_singleton, toWord] - exact reduce_replicate _ _ - -@[to_additive (attr := simp)] -theorem toWord_eq_nil_iff {x : FreeGroup α} : x.toWord = [] ↔ x = 1 := - toWord_injective.eq_iff' toWord_one - -@[to_additive] -theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (reduce w) := by - apply reduce.min - rw [← red_invRev_iff, invRev_invRev] - apply Red.reduce_left - have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev - rwa [invRev_invRev] at this - -@[to_additive (attr := simp)] -theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by - rcases x with ⟨L⟩ - rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] - -@[to_additive] -lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by - refine Red.sublist ?_ - have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by - rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] - rw [this] - exact FreeGroup.reduce.red - -/-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ -@[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] -def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := - ⟨reduce L₁, reduce.rev H12, reduce.rev H13⟩ - -@[to_additive] -instance : DecidableEq (FreeGroup α) := - toWord_injective.decidableEq - --- TODO @[to_additive] doesn't succeed, possibly due to a bug -instance Red.decidableRel : DecidableRel (@Red α) - | [], [] => isTrue Red.refl - | [], _hd2 :: _tl2 => isFalse fun H => List.noConfusion (Red.nil_iff.1 H) - | (x, b) :: tl, [] => - match Red.decidableRel tl [(x, not b)] with - | isTrue H => isTrue <| Red.trans (Red.cons_cons H) <| (@Red.Step.not _ [] [] _ _).to_red - | isFalse H => isFalse fun H2 => H <| Red.cons_nil_iff_singleton.1 H2 - | (x1, b1) :: tl1, (x2, b2) :: tl2 => - if h : (x1, b1) = (x2, b2) then - match Red.decidableRel tl1 tl2 with - | isTrue H => isTrue <| h ▸ Red.cons_cons H - | isFalse H => isFalse fun H2 => H <| (Red.cons_cons_iff _).1 <| h.symm ▸ H2 - else - match Red.decidableRel tl1 ((x1, ! b1) :: (x2, b2) :: tl2) with - | isTrue H => isTrue <| (Red.cons_cons H).tail Red.Step.cons_not - | isFalse H => isFalse fun H2 => H <| Red.inv_of_red_of_ne h H2 - -/-- A list containing every word that `w₁` reduces to. -/ -def Red.enum (L₁ : List (α × Bool)) : List (List (α × Bool)) := - List.filter (Red L₁) (List.sublists L₁) - -theorem Red.enum.sound (H : L₂ ∈ List.filter (Red L₁) (List.sublists L₁)) : Red L₁ L₂ := - of_decide_eq_true (@List.of_mem_filter _ _ L₂ _ H) - -theorem Red.enum.complete (H : Red L₁ L₂) : L₂ ∈ Red.enum L₁ := - List.mem_filter_of_mem (List.mem_sublists.2 <| Red.sublist H) (decide_eq_true H) - -instance : Fintype { L₂ // Red L₁ L₂ } := - Fintype.subtype (List.toFinset <| Red.enum L₁) fun _L₂ => - ⟨fun H => Red.enum.sound <| List.mem_toFinset.1 H, fun H => - List.mem_toFinset.2 <| Red.enum.complete H⟩ - -end Reduce - -@[to_additive (attr := simp)] -theorem one_ne_of (a : α) : 1 ≠ of a := - letI := Classical.decEq α; ne_of_apply_ne toWord <| by simp - -@[to_additive (attr := simp)] -theorem of_ne_one (a : α) : of a ≠ 1 := one_ne_of _ |>.symm - -@[to_additive] -instance [Nonempty α] : Nontrivial (FreeGroup α) where - exists_pair_ne := let ⟨x⟩ := ‹Nonempty α›; ⟨1, of x, one_ne_of x⟩ - -section Metric - -variable [DecidableEq α] - -/-- The length of reduced words provides a norm on a free group. -/ -@[to_additive "The length of reduced words provides a norm on an additive free group."] -def norm (x : FreeGroup α) : ℕ := - x.toWord.length - -@[to_additive (attr := simp)] -theorem norm_inv_eq {x : FreeGroup α} : norm x⁻¹ = norm x := by - simp only [norm, toWord_inv, invRev_length] - -@[to_additive (attr := simp)] -theorem norm_eq_zero {x : FreeGroup α} : norm x = 0 ↔ x = 1 := by - simp only [norm, List.length_eq_zero, toWord_eq_nil_iff] - -@[to_additive (attr := simp)] -theorem norm_one : norm (1 : FreeGroup α) = 0 := - rfl - -@[to_additive (attr := simp)] -theorem norm_of (a : α) : norm (of a) = 1 := - rfl - -@[to_additive] -theorem norm_mk_le : norm (mk L₁) ≤ L₁.length := - reduce.red.length_le - -@[to_additive] -theorem norm_mul_le (x y : FreeGroup α) : norm (x * y) ≤ norm x + norm y := - calc - norm (x * y) = norm (mk (x.toWord ++ y.toWord)) := by rw [← mul_mk, mk_toWord, mk_toWord] - _ ≤ (x.toWord ++ y.toWord).length := norm_mk_le - _ = norm x + norm y := List.length_append _ _ - -@[to_additive (attr := simp)] -theorem norm_of_pow (a : α) (n : ℕ) : norm (of a ^ n) = n := by - rw [norm, toWord_of_pow, List.length_replicate] - -@[to_additive] -theorem norm_surjective [Nonempty α] : Function.Surjective (norm (α := α)) := by - let ⟨a⟩ := ‹Nonempty α› - exact Function.RightInverse.surjective <| norm_of_pow a - -end Metric - end FreeGroup diff --git a/Mathlib/GroupTheory/FreeGroup/Reduce.lean b/Mathlib/GroupTheory/FreeGroup/Reduce.lean new file mode 100644 index 0000000000000..3017c17281dd2 --- /dev/null +++ b/Mathlib/GroupTheory/FreeGroup/Reduce.lean @@ -0,0 +1,351 @@ +/- +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.GroupTheory.FreeGroup.Basic +import Mathlib.Data.Fintype.Basic +import Mathlib.Data.List.Sublists + +/-! +# The maximal reduction of a word in a free group + +## Main declarations + +* `FreeGroup.reduce`: the maximal reduction of a word in a free group +* `FreeGroup.norm`: the length of the maximal reduction of a word in a free group + +-/ + + +namespace FreeGroup + +variable {α : Type*} +variable {L L₁ L₂ L₃ L₄ : List (α × Bool)} + +section Reduce + +variable [DecidableEq α] + +/-- The maximal reduction of a word. It is computable +iff `α` has decidable equality. -/ +@[to_additive "The maximal reduction of a word. It is computable iff `α` has decidable equality."] +def reduce : (L : List (α × Bool)) -> List (α × Bool) := + List.rec [] fun hd1 _tl1 ih => + List.casesOn ih [hd1] fun hd2 tl2 => + if hd1.1 = hd2.1 ∧ hd1.2 = not hd2.2 then tl2 else hd1 :: hd2 :: tl2 + +@[to_additive (attr := simp)] lemma reduce_nil : reduce ([] : List (α × Bool)) = [] := rfl +@[to_additive] lemma reduce_singleton (s : α × Bool) : reduce [s] = [s] := rfl + +@[to_additive (attr := simp)] +theorem reduce.cons (x) : + reduce (x :: L) = + List.casesOn (reduce L) [x] fun hd tl => + if x.1 = hd.1 ∧ x.2 = not hd.2 then tl else x :: hd :: tl := + rfl + +@[to_additive (attr := simp)] +theorem reduce_replicate (n : ℕ) (x : α × Bool) : + reduce (.replicate n x) = .replicate n x := by + induction n with + | zero => simp [reduce] + | succ n ih => + rw [List.replicate_succ, reduce.cons, ih] + cases n with + | zero => simp + | succ n => simp [List.replicate_succ] + +/-- The first theorem that characterises the function `reduce`: a word reduces to its maximal + reduction. -/ +@[to_additive "The first theorem that characterises the function `reduce`: a word reduces to its + maximal reduction."] +theorem reduce.red : Red L (reduce L) := by + induction L with + | nil => constructor + | cons hd1 tl1 ih => + dsimp + revert ih + generalize htl : reduce tl1 = TL + intro ih + cases TL with + | nil => exact Red.cons_cons ih + | cons hd2 tl2 => + dsimp only + split_ifs with h + · cases hd1 + cases hd2 + cases h + dsimp at * + subst_vars + apply Red.trans (Red.cons_cons ih) + exact Red.Step.cons_not_rev.to_red + · exact Red.cons_cons ih + +@[to_additive] +theorem reduce.not {p : Prop} : + ∀ {L₁ L₂ L₃ : List (α × Bool)} {x b}, reduce L₁ = L₂ ++ (x, b) :: (x, !b) :: L₃ → p + | [], L2, L3, _, _ => fun h => by cases L2 <;> injections + | (x, b) :: L1, L2, L3, x', b' => by + dsimp + cases r : reduce L1 with + | nil => + dsimp; intro h + exfalso + have := congr_arg List.length h + simp? [List.length] at this says + simp only [List.length, zero_add, List.length_append] at this + rw [add_comm, add_assoc, add_assoc, add_comm, <-add_assoc] at this + omega + | cons hd tail => + cases' hd with y c + dsimp only + split_ifs with h <;> intro H + · rw [H] at r + exact @reduce.not _ L1 ((y, c) :: L2) L3 x' b' r + rcases L2 with (_ | ⟨a, L2⟩) + · injections; subst_vars + simp at h + · refine @reduce.not _ L1 L2 L3 x' b' ?_ + injection H with _ H + rw [r, H]; rfl + +/-- The second theorem that characterises the function `reduce`: the maximal reduction of a word +only reduces to itself. -/ +@[to_additive "The second theorem that characterises the function `reduce`: the maximal reduction of + a word only reduces to itself."] +theorem reduce.min (H : Red (reduce L₁) L₂) : reduce L₁ = L₂ := by + induction' H with L1 L' L2 H1 H2 ih + · rfl + · cases' H1 with L4 L5 x b + exact reduce.not H2 + +/-- `reduce` is idempotent, i.e. the maximal reduction of the maximal reduction of a word is the + maximal reduction of the word. -/ +@[to_additive (attr := simp) "`reduce` is idempotent, i.e. the maximal reduction of the maximal + reduction of a word is the maximal reduction of the word."] +theorem reduce.idem : reduce (reduce L) = reduce L := + Eq.symm <| reduce.min reduce.red + +@[to_additive] +theorem reduce.Step.eq (H : Red.Step L₁ L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (reduce.red.head H) + (reduce.min HR13).trans (reduce.min HR23).symm + +/-- If a word reduces to another word, then they have a common maximal reduction. -/ +@[to_additive "If a word reduces to another word, then they have a common maximal reduction."] +theorem reduce.eq_of_red (H : Red L₁ L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, HR13, HR23⟩ := Red.church_rosser reduce.red (Red.trans H reduce.red) + (reduce.min HR13).trans (reduce.min HR23).symm + +alias red.reduce_eq := reduce.eq_of_red + +alias freeAddGroup.red.reduce_eq := FreeAddGroup.reduce.eq_of_red + +@[to_additive] +theorem Red.reduce_right (h : Red L₁ L₂) : Red L₁ (reduce L₂) := + reduce.eq_of_red h ▸ reduce.red + +@[to_additive] +theorem Red.reduce_left (h : Red L₁ L₂) : Red L₂ (reduce L₁) := + (reduce.eq_of_red h).symm ▸ reduce.red + +/-- If two words correspond to the same element in the free group, then they +have a common maximal reduction. This is the proof that the function that sends +an element of the free group to its maximal reduction is well-defined. -/ +@[to_additive "If two words correspond to the same element in the additive free group, then they + have a common maximal reduction. This is the proof that the function that sends an element of the + free group to its maximal reduction is well-defined."] +theorem reduce.sound (H : mk L₁ = mk L₂) : reduce L₁ = reduce L₂ := + let ⟨_L₃, H13, H23⟩ := Red.exact.1 H + (reduce.eq_of_red H13).trans (reduce.eq_of_red H23).symm + +/-- If two words have a common maximal reduction, then they correspond to the same element in the + free group. -/ +@[to_additive "If two words have a common maximal reduction, then they correspond to the same + element in the additive free group."] +theorem reduce.exact (H : reduce L₁ = reduce L₂) : mk L₁ = mk L₂ := + Red.exact.2 ⟨reduce L₂, H ▸ reduce.red, reduce.red⟩ + +/-- A word and its maximal reduction correspond to the same element of the free group. -/ +@[to_additive "A word and its maximal reduction correspond to the same element of the additive free + group."] +theorem reduce.self : mk (reduce L) = mk L := + reduce.exact reduce.idem + +/-- If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal reduction + of `w₁`. -/ +@[to_additive "If words `w₁ w₂` are such that `w₁` reduces to `w₂`, then `w₂` reduces to the maximal + reduction of `w₁`."] +theorem reduce.rev (H : Red L₁ L₂) : Red L₂ (reduce L₁) := + (reduce.eq_of_red H).symm ▸ reduce.red + +/-- The function that sends an element of the free group to its maximal reduction. -/ +@[to_additive "The function that sends an element of the additive free group to its maximal + reduction."] +def toWord : FreeGroup α → List (α × Bool) := + Quot.lift reduce fun _L₁ _L₂ H => reduce.Step.eq H + +@[to_additive] +theorem mk_toWord : ∀ {x : FreeGroup α}, mk (toWord x) = x := by rintro ⟨L⟩; exact reduce.self + +@[to_additive] +theorem toWord_injective : Function.Injective (toWord : FreeGroup α → List (α × Bool)) := by + rintro ⟨L₁⟩ ⟨L₂⟩; exact reduce.exact + +@[to_additive (attr := simp)] +theorem toWord_inj {x y : FreeGroup α} : toWord x = toWord y ↔ x = y := + toWord_injective.eq_iff + +@[to_additive (attr := simp)] +theorem toWord_mk : (mk L₁).toWord = reduce L₁ := + rfl + +@[to_additive (attr := simp)] +theorem toWord_of (a : α) : (of a).toWord = [(a, true)] := + rfl + +@[to_additive (attr := simp)] +theorem reduce_toWord : ∀ x : FreeGroup α, reduce (toWord x) = toWord x := by + rintro ⟨L⟩ + exact reduce.idem + +@[to_additive (attr := simp)] +theorem toWord_one : (1 : FreeGroup α).toWord = [] := + rfl + +@[to_additive (attr := simp)] +theorem toWord_of_pow (a : α) (n : ℕ) : (of a ^ n).toWord = List.replicate n (a, true) := by + rw [of, pow_mk, List.flatten_replicate_singleton, toWord] + exact reduce_replicate _ _ + +@[to_additive (attr := simp)] +theorem toWord_eq_nil_iff {x : FreeGroup α} : x.toWord = [] ↔ x = 1 := + toWord_injective.eq_iff' toWord_one + +@[to_additive] +theorem reduce_invRev {w : List (α × Bool)} : reduce (invRev w) = invRev (reduce w) := by + apply reduce.min + rw [← red_invRev_iff, invRev_invRev] + apply Red.reduce_left + have : Red (invRev (invRev w)) (invRev (reduce (invRev w))) := reduce.red.invRev + rwa [invRev_invRev] at this + +@[to_additive (attr := simp)] +theorem toWord_inv (x : FreeGroup α) : x⁻¹.toWord = invRev x.toWord := by + rcases x with ⟨L⟩ + rw [quot_mk_eq_mk, inv_mk, toWord_mk, toWord_mk, reduce_invRev] + +open List -- for <+ notation + +@[to_additive] +lemma toWord_mul_sublist (x y : FreeGroup α) : (x * y).toWord <+ x.toWord ++ y.toWord := by + refine Red.sublist ?_ + have : x * y = FreeGroup.mk (x.toWord ++ y.toWord) := by + rw [← FreeGroup.mul_mk, FreeGroup.mk_toWord, FreeGroup.mk_toWord] + rw [this] + exact FreeGroup.reduce.red + +/-- **Constructive Church-Rosser theorem** (compare `church_rosser`). -/ +@[to_additive "**Constructive Church-Rosser theorem** (compare `church_rosser`)."] +def reduce.churchRosser (H12 : Red L₁ L₂) (H13 : Red L₁ L₃) : { L₄ // Red L₂ L₄ ∧ Red L₃ L₄ } := + ⟨reduce L₁, reduce.rev H12, reduce.rev H13⟩ + +@[to_additive] +instance : DecidableEq (FreeGroup α) := + toWord_injective.decidableEq + +-- TODO @[to_additive] doesn't succeed, possibly due to a bug +instance Red.decidableRel : DecidableRel (@Red α) + | [], [] => isTrue Red.refl + | [], _hd2 :: _tl2 => isFalse fun H => List.noConfusion (Red.nil_iff.1 H) + | (x, b) :: tl, [] => + match Red.decidableRel tl [(x, not b)] with + | isTrue H => isTrue <| Red.trans (Red.cons_cons H) <| (@Red.Step.not _ [] [] _ _).to_red + | isFalse H => isFalse fun H2 => H <| Red.cons_nil_iff_singleton.1 H2 + | (x1, b1) :: tl1, (x2, b2) :: tl2 => + if h : (x1, b1) = (x2, b2) then + match Red.decidableRel tl1 tl2 with + | isTrue H => isTrue <| h ▸ Red.cons_cons H + | isFalse H => isFalse fun H2 => H <| (Red.cons_cons_iff _).1 <| h.symm ▸ H2 + else + match Red.decidableRel tl1 ((x1, ! b1) :: (x2, b2) :: tl2) with + | isTrue H => isTrue <| (Red.cons_cons H).tail Red.Step.cons_not + | isFalse H => isFalse fun H2 => H <| Red.inv_of_red_of_ne h H2 + +/-- A list containing every word that `w₁` reduces to. -/ +def Red.enum (L₁ : List (α × Bool)) : List (List (α × Bool)) := + List.filter (Red L₁) (List.sublists L₁) + +theorem Red.enum.sound (H : L₂ ∈ List.filter (Red L₁) (List.sublists L₁)) : Red L₁ L₂ := + of_decide_eq_true (@List.of_mem_filter _ _ L₂ _ H) + +theorem Red.enum.complete (H : Red L₁ L₂) : L₂ ∈ Red.enum L₁ := + List.mem_filter_of_mem (List.mem_sublists.2 <| Red.sublist H) (decide_eq_true H) + +instance (L₁ : List (α × Bool)) : Fintype { L₂ // Red L₁ L₂ } := + Fintype.subtype (List.toFinset <| Red.enum L₁) fun _L₂ => + ⟨fun H => Red.enum.sound <| List.mem_toFinset.1 H, fun H => + List.mem_toFinset.2 <| Red.enum.complete H⟩ + +end Reduce + +@[to_additive (attr := simp)] +theorem one_ne_of (a : α) : 1 ≠ of a := + letI := Classical.decEq α; ne_of_apply_ne toWord <| by simp + +@[to_additive (attr := simp)] +theorem of_ne_one (a : α) : of a ≠ 1 := one_ne_of _ |>.symm + +@[to_additive] +instance [Nonempty α] : Nontrivial (FreeGroup α) where + exists_pair_ne := let ⟨x⟩ := ‹Nonempty α›; ⟨1, of x, one_ne_of x⟩ + +section Metric + +variable [DecidableEq α] + +/-- The length of reduced words provides a norm on a free group. -/ +@[to_additive "The length of reduced words provides a norm on an additive free group."] +def norm (x : FreeGroup α) : ℕ := + x.toWord.length + +@[to_additive (attr := simp)] +theorem norm_inv_eq {x : FreeGroup α} : norm x⁻¹ = norm x := by + simp only [norm, toWord_inv, invRev_length] + +@[to_additive (attr := simp)] +theorem norm_eq_zero {x : FreeGroup α} : norm x = 0 ↔ x = 1 := by + simp only [norm, List.length_eq_zero, toWord_eq_nil_iff] + +@[to_additive (attr := simp)] +theorem norm_one : norm (1 : FreeGroup α) = 0 := + rfl + +@[to_additive (attr := simp)] +theorem norm_of (a : α) : norm (of a) = 1 := + rfl + +@[to_additive] +theorem norm_mk_le : norm (mk L₁) ≤ L₁.length := + reduce.red.length_le + +@[to_additive] +theorem norm_mul_le (x y : FreeGroup α) : norm (x * y) ≤ norm x + norm y := + calc + norm (x * y) = norm (mk (x.toWord ++ y.toWord)) := by rw [← mul_mk, mk_toWord, mk_toWord] + _ ≤ (x.toWord ++ y.toWord).length := norm_mk_le + _ = norm x + norm y := List.length_append _ _ + +@[to_additive (attr := simp)] +theorem norm_of_pow (a : α) (n : ℕ) : norm (of a ^ n) = n := by + rw [norm, toWord_of_pow, List.length_replicate] + +@[to_additive] +theorem norm_surjective [Nonempty α] : Function.Surjective (norm (α := α)) := by + let ⟨a⟩ := ‹Nonempty α› + exact Function.RightInverse.surjective <| norm_of_pow a + +end Metric + +end FreeGroup diff --git a/Mathlib/SetTheory/Cardinal/Free.lean b/Mathlib/SetTheory/Cardinal/Free.lean index d802c1a3f9335..074fafe8ae1f3 100644 --- a/Mathlib/SetTheory/Cardinal/Free.lean +++ b/Mathlib/SetTheory/Cardinal/Free.lean @@ -5,6 +5,7 @@ Authors: Eric Wieser, Daniel Weber -/ import Mathlib.Data.Finsupp.Fintype import Mathlib.GroupTheory.FreeAbelianGroupFinsupp +import Mathlib.GroupTheory.FreeGroup.Reduce import Mathlib.RingTheory.FreeCommRing import Mathlib.SetTheory.Cardinal.Arithmetic import Mathlib.SetTheory.Cardinal.Finsupp From ebd1b1af5c41902abbe4beb77dc8c7db0c32f209 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 09:56:45 +0000 Subject: [PATCH 218/250] =?UTF-8?q?feat(Algebra):=20ULift=20=E2=84=A4=20is?= =?UTF-8?q?=20initial=20in=20the=20category=20of=20commutative=20rings=20(?= =?UTF-8?q?#19507)?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Algebra/Category/Ring/Constructions.lean | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/Mathlib/Algebra/Category/Ring/Constructions.lean b/Mathlib/Algebra/Category/Ring/Constructions.lean index ca21d51022790..c124625d6a67f 100644 --- a/Mathlib/Algebra/Category/Ring/Constructions.lean +++ b/Mathlib/Algebra/Category/Ring/Constructions.lean @@ -144,6 +144,14 @@ theorem subsingleton_of_isTerminal {X : CommRingCat} (hX : IsTerminal X) : Subsi def zIsInitial : IsInitial (CommRingCat.of ℤ) := IsInitial.ofUnique (h := fun R => ⟨⟨Int.castRingHom R⟩, fun a => a.ext_int _⟩) +/-- `ULift.{u} ℤ` is initial in `CommRingCat`. -/ +def isInitial : IsInitial (CommRingCat.of (ULift.{u} ℤ)) := + IsInitial.ofUnique (h := fun R ↦ ⟨⟨(Int.castRingHom R).comp ULift.ringEquiv.toRingHom⟩, + fun _ ↦ by + rw [← RingHom.cancel_right (f := (ULift.ringEquiv.{0, u} (α := ℤ)).symm.toRingHom) + (hf := ULift.ringEquiv.symm.surjective)] + apply RingHom.ext_int⟩) + end Terminal section Product From 3942daa810bd06acd4daefa3fc3771378c5b2acb Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 09:56:46 +0000 Subject: [PATCH 219/250] chore(CategoryTheory/Sites): split DenseSubsite (#19519) The file `CategoryTheory.Sites.DenseSubsite` is split into `DenseSubsite.Basic` and `DenseSubsite.SheafEquiv`, the latter containing the equivalence of categories of sheaves in case suitable limits exists. The file `Sites.InducedTopology` is also moved to `Sites.DenseSubsite`. (This split prepares for #19444 where the equivalence of categories of sheaves is obtained under a somewhat weaker condition.) --- Mathlib.lean | 5 +- .../Sites/Coherent/SheafComparison.lean | 2 +- .../CategoryTheory/Sites/ConstantSheaf.lean | 2 +- .../Basic.lean} | 123 --------------- .../{ => DenseSubsite}/InducedTopology.lean | 2 +- .../Sites/DenseSubsite/SheafEquiv.lean | 142 ++++++++++++++++++ Mathlib/CategoryTheory/Sites/Equivalence.lean | 2 +- .../Sites/PreservesLocallyBijective.lean | 2 +- .../Sheaves/SheafCondition/Sites.lean | 2 +- 9 files changed, 151 insertions(+), 131 deletions(-) rename Mathlib/CategoryTheory/Sites/{DenseSubsite.lean => DenseSubsite/Basic.lean} (78%) rename Mathlib/CategoryTheory/Sites/{ => DenseSubsite}/InducedTopology.lean (98%) create mode 100644 Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean diff --git a/Mathlib.lean b/Mathlib.lean index ff6ae6de6f371..521d0acd48e02 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -2073,13 +2073,14 @@ import Mathlib.CategoryTheory.Sites.CoverLifting import Mathlib.CategoryTheory.Sites.CoverPreserving import Mathlib.CategoryTheory.Sites.Coverage import Mathlib.CategoryTheory.Sites.CoversTop -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv import Mathlib.CategoryTheory.Sites.EffectiveEpimorphic import Mathlib.CategoryTheory.Sites.EpiMono import Mathlib.CategoryTheory.Sites.EqualizerSheafCondition import Mathlib.CategoryTheory.Sites.Equivalence import Mathlib.CategoryTheory.Sites.Grothendieck -import Mathlib.CategoryTheory.Sites.InducedTopology import Mathlib.CategoryTheory.Sites.IsSheafFor import Mathlib.CategoryTheory.Sites.IsSheafOneHypercover import Mathlib.CategoryTheory.Sites.LeftExact diff --git a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean index e78d9e4dc5cb6..5137ac167f17e 100644 --- a/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean +++ b/Mathlib/CategoryTheory/Sites/Coherent/SheafComparison.lean @@ -7,7 +7,7 @@ import Mathlib.CategoryTheory.Sites.Coherent.Comparison import Mathlib.CategoryTheory.Sites.Coherent.ExtensiveSheaves import Mathlib.CategoryTheory.Sites.Coherent.ReflectsPrecoherent import Mathlib.CategoryTheory.Sites.Coherent.ReflectsPreregular -import Mathlib.CategoryTheory.Sites.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology import Mathlib.CategoryTheory.Sites.Whiskering /-! diff --git a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean index a300122560b4c..9bc3104f4e593 100644 --- a/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean +++ b/Mathlib/CategoryTheory/Sites/ConstantSheaf.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ import Mathlib.CategoryTheory.Sites.Sheafification -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv /-! # The constant sheaf diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean similarity index 78% rename from Mathlib/CategoryTheory/Sites/DenseSubsite.lean rename to Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean index 5e24ab6fe6ff6..a7b6c780ed15f 100644 --- a/Mathlib/CategoryTheory/Sites/DenseSubsite.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/Basic.lean @@ -29,9 +29,6 @@ that factors through images of the functor for each object in `D`. - `CategoryTheory.Functor.IsDenseSubsite`: The functor `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)` if `G` is cover-dense, locally fully-faithful, and `S` is a cover of `C` iff the image of `S` in `D` is a cover. -- `CategoryTheory.Functor.IsDenseSubsite.sheafEquiv`: - If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, - it induces an equivalence of category of sheaves valued in a category with suitable limits. ## References @@ -41,7 +38,6 @@ that factors through images of the functor for each object in `D`. -/ - universe w v u namespace CategoryTheory @@ -556,122 +552,3 @@ end IsDenseSubsite end Functor end CategoryTheory - -namespace CategoryTheory.Functor.IsDenseSubsite - -open CategoryTheory Opposite - -universe w' -variable {C D : Type*} [Category C] [Category D] -variable (G : C ⥤ D) -variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) -variable {A : Type w} [Category.{w'} A] [∀ X, Limits.HasLimitsOfShape (StructuredArrow X G.op) A] -variable [G.IsDenseSubsite J K] - -include K in -lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : - IsIso ((yoneda.map ((G.op.ranCounit.app Y.val).app (op U))).app (op X)) := by - rw [isIso_iff_bijective] - constructor - · intro f₁ f₂ e - apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext - rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ - obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ - apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext - dsimp - rintro V iVW ⟨iVU, e'⟩ - have := congr($e ≫ Y.1.map iVU.op) - simp only [comp_obj, yoneda_map_app, Category.assoc, coyoneda_obj_obj, comp_map, - coyoneda_obj_map, ← NatTrans.naturality, op_obj, op_map, Quiver.Hom.unop_op, ← map_comp_assoc, - ← op_comp, ← e'] at this ⊢ - erw [← NatTrans.naturality] at this - exact this - · intro f - have (X Y Z) (f : X ⟶ Y) (g : G.obj Y ⟶ G.obj Z) (hf : G.imageSieve g f) : Exists _ := hf - choose l hl using this - let c : Limits.Cone (StructuredArrow.proj (op (G.obj U)) G.op ⋙ Y.val) := by - refine ⟨X, ⟨fun g ↦ ?_, ?_⟩⟩ - · refine Y.2.amalgamate ⟨_, IsDenseSubsite.imageSieve_mem J K G g.hom.unop⟩ - (fun I ↦ f ≫ Y.1.map (l _ _ _ _ _ I.hf).op) fun I₁ I₂ r ↦ ?_ - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (r.g₁ ≫ l _ _ _ _ _ I₁.hf) - (r.g₂ ≫ l _ _ _ _ _ I₂.hf) ?_)).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - · simp only [const_obj_obj, op_obj, map_comp, hl] - simp only [← map_comp_assoc, r.w] - · simp [← map_comp, ← op_comp, hiUV] - · dsimp - rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ - dsimp at g₁ g₂ i hi - -- See issue https://github.com/leanprover-community/mathlib4/pull/15781 for tracking performance regressions of `rintro` as here - have h : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi - rcases h with ⟨rfl⟩ - have h : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ - rcases h with ⟨g, rfl⟩ - have h : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ - rcases h with ⟨i, rfl⟩ - simp only [const_obj_obj, id_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, op_obj, - unop_comp, Quiver.Hom.unop_op, Category.id_comp, comp_map, StructuredArrow.proj_map] - apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (G.map i ≫ g)⟩ - intro I - simp only [Presheaf.IsSheaf.amalgamate_map, Category.assoc, ← Functor.map_comp, ← op_comp] - let I' : GrothendieckTopology.Cover.Arrow ⟨_, IsDenseSubsite.imageSieve_mem J K G g⟩ := - ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ - refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) - (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - simp [← Functor.map_comp, ← op_comp, hiUV] - refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ - · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) - simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, - RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, - const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, - RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this - simp only [id_obj, yoneda_map_app, this] - apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ - apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) - I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ - simp [← Functor.map_comp, ← op_comp, hiUV] - -instance (Y : Sheaf J A) : IsIso ((G.sheafAdjunctionCocontinuous A J K).counit.app Y) := by - apply (config := { allowSynthFailures := true }) - ReflectsIsomorphisms.reflects (sheafToPresheaf J A) - rw [NatTrans.isIso_iff_isIso_app] - intro ⟨U⟩ - apply (config := { allowSynthFailures := true }) ReflectsIsomorphisms.reflects yoneda - rw [NatTrans.isIso_iff_isIso_app] - intro ⟨X⟩ - simp only [comp_obj, sheafToPresheaf_obj, sheafPushforwardContinuous_obj_val_obj, yoneda_obj_obj, - id_obj, sheafToPresheaf_map, sheafAdjunctionCocontinuous_counit_app_val, ranAdjunction_counit] - exact isIso_ranCounit_app_of_isDenseSubsite G J K Y U X - -variable (A) - -/-- -If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, -it induces an equivalence of category of sheaves valued in a category with suitable limits. --/ -@[simps! functor inverse] -noncomputable def sheafEquiv : Sheaf J A ≌ Sheaf K A := - (G.sheafAdjunctionCocontinuous A J K).toEquivalence.symm - -instance : (G.sheafPushforwardContinuous A J K).IsEquivalence := - inferInstanceAs (IsDenseSubsite.sheafEquiv G _ _ _).inverse.IsEquivalence - -variable [HasWeakSheafify J A] [HasWeakSheafify K A] - -/-- The natural isomorphism exhibiting the compatibility of -`IsDenseSubsite.sheafEquiv` with sheafification. -/ -noncomputable -abbrev sheafEquivSheafificationCompatibility : - (whiskeringLeft _ _ A).obj G.op ⋙ presheafToSheaf _ _ ≅ - presheafToSheaf _ _ ⋙ (sheafEquiv G J K A).inverse := by - apply Functor.pushforwardContinuousSheafificationCompatibility - -end IsDenseSubsite - -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv -@[deprecated (since := "2024-07-23")] -alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := - IsDenseSubsite.sheafEquivSheafificationCompatibility - -end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/InducedTopology.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean similarity index 98% rename from Mathlib/CategoryTheory/Sites/InducedTopology.lean rename to Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean index c5be7f7e1ffda..4710ae5d4bf61 100644 --- a/Mathlib/CategoryTheory/Sites/InducedTopology.lean +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/InducedTopology.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Andrew Yang. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Andrew Yang -/ -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.SheafEquiv /-! # Induced Topology diff --git a/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean new file mode 100644 index 0000000000000..7cc5b9aefa271 --- /dev/null +++ b/Mathlib/CategoryTheory/Sites/DenseSubsite/SheafEquiv.lean @@ -0,0 +1,142 @@ +/- +Copyright (c) 2021 Andrew Yang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Andrew Yang +-/ + +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic + +/-! +# The equivalence of categories of sheaves of a dense subsite + +- `CategoryTheory.Functor.IsDenseSubsite.sheafEquiv`: + If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, + it induces an equivalence of category of sheaves valued in a category with suitable limits. + +## References + +* [Elephant]: *Sketches of an Elephant*, ℱ. T. Johnstone: C2.2. +* https://ncatlab.org/nlab/show/dense+sub-site +* https://ncatlab.org/nlab/show/comparison+lemma + +-/ + +universe w v u w' + +namespace CategoryTheory.Functor.IsDenseSubsite + +open CategoryTheory Opposite + +variable {C D : Type*} [Category C] [Category D] +variable (G : C ⥤ D) +variable (J : GrothendieckTopology C) (K : GrothendieckTopology D) +variable {A : Type w} [Category.{w'} A] [∀ X, Limits.HasLimitsOfShape (StructuredArrow X G.op) A] +variable [G.IsDenseSubsite J K] + +include K in +lemma isIso_ranCounit_app_of_isDenseSubsite (Y : Sheaf J A) (U X) : + IsIso ((yoneda.map ((G.op.ranCounit.app Y.val).app (op U))).app (op X)) := by + rw [isIso_iff_bijective] + constructor + · intro f₁ f₂ e + apply (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).hom_ext + rintro ⟨⟨⟨⟩⟩, ⟨W⟩, g⟩ + obtain ⟨g, rfl⟩ : ∃ g' : G.obj W ⟶ G.obj U, g = g'.op := ⟨g.unop, rfl⟩ + apply (Y.2 X _ (IsDenseSubsite.imageSieve_mem J K G g)).isSeparatedFor.ext + dsimp + rintro V iVW ⟨iVU, e'⟩ + have := congr($e ≫ Y.1.map iVU.op) + simp only [comp_obj, yoneda_map_app, Category.assoc, coyoneda_obj_obj, comp_map, + coyoneda_obj_map, ← NatTrans.naturality, op_obj, op_map, Quiver.Hom.unop_op, ← map_comp_assoc, + ← op_comp, ← e'] at this ⊢ + erw [← NatTrans.naturality] at this + exact this + · intro f + have (X Y Z) (f : X ⟶ Y) (g : G.obj Y ⟶ G.obj Z) (hf : G.imageSieve g f) : Exists _ := hf + choose l hl using this + let c : Limits.Cone (StructuredArrow.proj (op (G.obj U)) G.op ⋙ Y.val) := by + refine ⟨X, ⟨fun g ↦ ?_, ?_⟩⟩ + · refine Y.2.amalgamate ⟨_, IsDenseSubsite.imageSieve_mem J K G g.hom.unop⟩ + (fun I ↦ f ≫ Y.1.map (l _ _ _ _ _ I.hf).op) fun I₁ I₂ r ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (r.g₁ ≫ l _ _ _ _ _ I₁.hf) + (r.g₂ ≫ l _ _ _ _ _ I₂.hf) ?_)).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + · simp only [const_obj_obj, op_obj, map_comp, hl] + simp only [← map_comp_assoc, r.w] + · simp [← map_comp, ← op_comp, hiUV] + · dsimp + rintro ⟨⟨⟨⟩⟩, ⟨W₁⟩, g₁⟩ ⟨⟨⟨⟩⟩, ⟨W₂⟩, g₂⟩ ⟨⟨⟨⟨⟩⟩⟩, i, hi⟩ + dsimp at g₁ g₂ i hi + -- See issue https://github.com/leanprover-community/mathlib4/pull/15781 for tracking performance regressions of `rintro` as here + have h : g₂ = g₁ ≫ (G.map i.unop).op := by simpa only [Category.id_comp] using hi + rcases h with ⟨rfl⟩ + have h : ∃ g' : G.obj W₁ ⟶ G.obj U, g₁ = g'.op := ⟨g₁.unop, rfl⟩ + rcases h with ⟨g, rfl⟩ + have h : ∃ i' : W₂ ⟶ W₁, i = i'.op := ⟨i.unop, rfl⟩ + rcases h with ⟨i, rfl⟩ + simp only [const_obj_obj, id_obj, comp_obj, StructuredArrow.proj_obj, const_obj_map, op_obj, + unop_comp, Quiver.Hom.unop_op, Category.id_comp, comp_map, StructuredArrow.proj_map] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (G.map i ≫ g)⟩ + intro I + simp only [Presheaf.IsSheaf.amalgamate_map, Category.assoc, ← Functor.map_comp, ← op_comp] + let I' : GrothendieckTopology.Cover.Arrow ⟨_, IsDenseSubsite.imageSieve_mem J K G g⟩ := + ⟨_, I.f ≫ i, ⟨l _ _ _ _ _ I.hf, by simp [hl]⟩⟩ + refine Eq.trans ?_ (Y.2.amalgamate_map _ _ _ I').symm + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + (l _ _ _ _ _ I'.hf) (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] + refine ⟨(isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).lift c, ?_⟩ + · have := (isPointwiseRightKanExtensionRanCounit G.op Y.1 (.op (G.obj U))).fac c (.mk (𝟙 _)) + simp only [id_obj, comp_obj, StructuredArrow.proj_obj, StructuredArrow.mk_right, + RightExtension.coneAt_pt, RightExtension.mk_left, RightExtension.coneAt_π_app, + const_obj_obj, op_obj, StructuredArrow.mk_hom_eq_self, map_id, whiskeringLeft_obj_obj, + RightExtension.mk_hom, Category.id_comp, StructuredArrow.mk_left, unop_id] at this + simp only [id_obj, yoneda_map_app, this] + apply Y.2.hom_ext ⟨_, IsDenseSubsite.imageSieve_mem J K G (𝟙 (G.obj U))⟩ _ _ fun I ↦ ?_ + apply (Y.2 X _ (IsDenseSubsite.equalizer_mem J K G (l _ _ _ _ _ I.hf) + I.f (by simp [hl]))).isSeparatedFor.ext fun V iUV (hiUV : _ = _) ↦ ?_ + simp [← Functor.map_comp, ← op_comp, hiUV] + +instance (Y : Sheaf J A) : IsIso ((G.sheafAdjunctionCocontinuous A J K).counit.app Y) := by + apply (config := { allowSynthFailures := true }) + ReflectsIsomorphisms.reflects (sheafToPresheaf J A) + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨U⟩ + apply (config := { allowSynthFailures := true }) ReflectsIsomorphisms.reflects yoneda + rw [NatTrans.isIso_iff_isIso_app] + intro ⟨X⟩ + simp only [comp_obj, sheafToPresheaf_obj, sheafPushforwardContinuous_obj_val_obj, yoneda_obj_obj, + id_obj, sheafToPresheaf_map, sheafAdjunctionCocontinuous_counit_app_val, ranAdjunction_counit] + exact isIso_ranCounit_app_of_isDenseSubsite G J K Y U X + +variable (A) + +/-- +If `G : C ⥤ D` exhibits `(C, J)` as a dense subsite of `(D, K)`, +it induces an equivalence of category of sheaves valued in a category with suitable limits. +-/ +@[simps! functor inverse] +noncomputable def sheafEquiv : Sheaf J A ≌ Sheaf K A := + (G.sheafAdjunctionCocontinuous A J K).toEquivalence.symm + +instance : (G.sheafPushforwardContinuous A J K).IsEquivalence := + inferInstanceAs (IsDenseSubsite.sheafEquiv G _ _ _).inverse.IsEquivalence + +variable [HasWeakSheafify J A] [HasWeakSheafify K A] + +/-- The natural isomorphism exhibiting the compatibility of +`IsDenseSubsite.sheafEquiv` with sheafification. -/ +noncomputable +abbrev sheafEquivSheafificationCompatibility : + (whiskeringLeft _ _ A).obj G.op ⋙ presheafToSheaf _ _ ≅ + presheafToSheaf _ _ ⋙ (sheafEquiv G J K A).inverse := by + apply Functor.pushforwardContinuousSheafificationCompatibility + +end IsDenseSubsite + +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLifting := IsDenseSubsite.sheafEquiv +@[deprecated (since := "2024-07-23")] +alias IsCoverDense.sheafEquivOfCoverPreservingCoverLiftingSheafificationCompatibility := + IsDenseSubsite.sheafEquivSheafificationCompatibility + +end CategoryTheory.Functor diff --git a/Mathlib/CategoryTheory/Sites/Equivalence.lean b/Mathlib/CategoryTheory/Sites/Equivalence.lean index 6953348d2088a..9ae28dadabe3f 100644 --- a/Mathlib/CategoryTheory/Sites/Equivalence.lean +++ b/Mathlib/CategoryTheory/Sites/Equivalence.lean @@ -3,7 +3,7 @@ Copyright (c) 2023 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.Sites.InducedTopology +import Mathlib.CategoryTheory.Sites.DenseSubsite.InducedTopology import Mathlib.CategoryTheory.Sites.LocallyBijective import Mathlib.CategoryTheory.Sites.PreservesLocallyBijective import Mathlib.CategoryTheory.Sites.Whiskering diff --git a/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean b/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean index bbc801acd097d..ab499ecc7dfde 100644 --- a/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean +++ b/Mathlib/CategoryTheory/Sites/PreservesLocallyBijective.lean @@ -3,7 +3,7 @@ Copyright (c) 2024 Dagur Asgeirsson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Dagur Asgeirsson -/ -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic import Mathlib.CategoryTheory.Sites.LocallySurjective /-! diff --git a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean index d4c904a20ecf8..185316bdf4310 100644 --- a/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean +++ b/Mathlib/Topology/Sheaves/SheafCondition/Sites.lean @@ -5,7 +5,7 @@ Authors: Justus Springer -/ import Mathlib.CategoryTheory.Sites.Spaces import Mathlib.Topology.Sheaves.Sheaf -import Mathlib.CategoryTheory.Sites.DenseSubsite +import Mathlib.CategoryTheory.Sites.DenseSubsite.Basic /-! From cfef5c4f0a57295f55f26a41cf24ae4d5ca2b613 Mon Sep 17 00:00:00 2001 From: Riccardo Brasca Date: Wed, 27 Nov 2024 10:10:36 +0000 Subject: [PATCH 220/250] doc: clarify doc of HasProd (#19530) --- Mathlib/Topology/Algebra/InfiniteSum/Defs.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean index cc83507487809..44ab58cf88b01 100644 --- a/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean +++ b/Mathlib/Topology/Algebra/InfiniteSum/Defs.lean @@ -59,7 +59,8 @@ section HasProd variable [CommMonoid α] [TopologicalSpace α] -/-- Infinite product on a topological monoid +/-- `HasProd f a` means that the (potentially infinite) product of the `f b` for `b : β` converges +to `a`. The `atTop` filter on `Finset β` is the limit of all finite sets towards the entire type. So we take the product over bigger and bigger sets. This product operation is invariant under reordering. @@ -70,7 +71,8 @@ this assumption later, for the lemmas where it is relevant. These are defined in an identical way to infinite sums (`HasSum`). For example, we say that the function `ℕ → ℝ` sending `n` to `1 / 2` has a product of `0`, rather than saying that it does not converge as some authors would. -/ -@[to_additive "Infinite sum on a topological monoid +@[to_additive "`HasSum f a` means that the (potentially infinite) sum of the `f b` for `b : β` +converges to `a`. The `atTop` filter on `Finset β` is the limit of all finite sets towards the entire type. So we sum up bigger and bigger sets. This sum operation is invariant under reordering. In particular, From c59545b6f988c2439d89047e58cef992b0dfa62a Mon Sep 17 00:00:00 2001 From: Markus Himmel Date: Wed, 27 Nov 2024 10:31:12 +0000 Subject: [PATCH 221/250] feat: ind-objects are closed under products (#18988) Co-authored-by: Markus Himmel --- Mathlib.lean | 2 + .../FilteredColimitCommutesProduct.lean | 185 ++++++++++++++++++ .../Limits/Indization/Products.lean | 43 ++++ 3 files changed, 230 insertions(+) create mode 100644 Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean create mode 100644 Mathlib/CategoryTheory/Limits/Indization/Products.lean diff --git a/Mathlib.lean b/Mathlib.lean index 521d0acd48e02..f3e568ce173a2 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1797,6 +1797,7 @@ import Mathlib.CategoryTheory.Limits.EssentiallySmall import Mathlib.CategoryTheory.Limits.ExactFunctor import Mathlib.CategoryTheory.Limits.Filtered import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesFiniteLimit +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesProduct import Mathlib.CategoryTheory.Limits.Final import Mathlib.CategoryTheory.Limits.Final.ParallelPair import Mathlib.CategoryTheory.Limits.FinallySmall @@ -1816,6 +1817,7 @@ import Mathlib.CategoryTheory.Limits.Indization.Equalizers import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits import Mathlib.CategoryTheory.Limits.Indization.IndObject import Mathlib.CategoryTheory.Limits.Indization.LocallySmall +import Mathlib.CategoryTheory.Limits.Indization.Products import Mathlib.CategoryTheory.Limits.IsConnected import Mathlib.CategoryTheory.Limits.IsLimit import Mathlib.CategoryTheory.Limits.Lattice diff --git a/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean new file mode 100644 index 0000000000000..ac195d5acac16 --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/FilteredColimitCommutesProduct.lean @@ -0,0 +1,185 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.FunctorCategory.Filtered +import Mathlib.CategoryTheory.Limits.Shapes.Types +import Mathlib.CategoryTheory.Limits.TypesFiltered +import Mathlib.CategoryTheory.Limits.FunctorCategory.Shapes.Products + +/-! +# The IPC property + +Given a family of categories `I i` (`i : α`) and a family of functors `F i : I i ⥤ C`, we construct +the natural morphism `colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)`. + +Similarly to the study of finite limits commuting with filtered colimits, we then study sufficient +conditions for this morphism to be an isomorphism. We say that `C` satisfies the `w`-IPC property if +the morphism is an isomorphism as long as `α` is `w`-small and `I i` is `w`-small and filtered for +all `i`. + +We show that +* the category `Type u` satisfies the `u`-IPC property and +* if `C` satisfies the `w`-IPC property, then `D ⥤ C` satisfies the `w`-IPC property. + +These results will be used to show that if a category `C` has products indexed by `α`, then so +does the category of Ind-objects of `C`. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], 3.1.10, 3.1.11, 3.1.12. +-/ + +universe w v v₁ v₂ u u₁ u₂ + +namespace CategoryTheory.Limits + +section + +variable {C : Type u} [Category.{v} C] {α : Type w} {I : α → Type u₁} [∀ i, Category.{v₁} (I i)] + [HasLimitsOfShape (Discrete α) C] + (F : ∀ i, I i ⥤ C) + +/-- Given a family of functors `I i ⥤ C` for `i : α`, we obtain a functor `(∀ i, I i) ⥤ C` which +maps `k : ∀ i, I i` to `∏ᶜ fun (s : α) => (F s).obj (k s)`. -/ +@[simps] +noncomputable def pointwiseProduct : (∀ i, I i) ⥤ C where + obj k := ∏ᶜ fun (s : α) => (F s).obj (k s) + map f := Pi.map (fun s => (F s).map (f s)) + +variable [∀ i, HasColimitsOfShape (I i) C] [HasColimitsOfShape (∀ i, I i) C] + +/-- The inclusions `(F s).obj (k s) ⟶ colimit (F s)` induce a cocone on `pointwiseProduct F` with +cone point `∏ᶜ (fun s : α) => colimit (F s)`. -/ +@[simps] +noncomputable def coconePointwiseProduct : Cocone (pointwiseProduct F) where + pt := ∏ᶜ fun (s : α) => colimit (F s) + ι := { app := fun k => Pi.map fun s => colimit.ι _ _ } + +/-- The natural morphism `colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)`. +We will say that a category has the `IPC` property if this morphism is an isomorphism as long +as the indexing categories are filtered. -/ +noncomputable def colimitPointwiseProductToProductColimit : + colimit (pointwiseProduct F) ⟶ ∏ᶜ fun (s : α) => colimit (F s) := + colimit.desc (pointwiseProduct F) (coconePointwiseProduct F) + +@[reassoc (attr := simp)] +theorem ι_colimitPointwiseProductToProductColimit_π (k : ∀ i, I i) (s : α) : + colimit.ι (pointwiseProduct F) k ≫ colimitPointwiseProductToProductColimit F ≫ Pi.π _ s = + Pi.π _ s ≫ colimit.ι (F s) (k s) := by + simp [colimitPointwiseProductToProductColimit] + +end + +section functorCategory + +variable {C : Type u} [Category.{v} C] {D : Type u₁} [Category.{v₁} D] + {α : Type w} {I : α → Type u₂} [∀ i, Category (I i)] + [HasLimitsOfShape (Discrete α) C] + (F : ∀ i, I i ⥤ D ⥤ C) + +/-- Evaluating the pointwise product `k ↦ ∏ᶜ fun (s : α) => (F s).obj (k s)` at `d` is the same as +taking the pointwise product `k ↦ ∏ᶜ fun (s : α) => ((F s).obj (k s)).obj d`. -/ +@[simps!] +noncomputable def pointwiseProductCompEvaluation (d : D) : + pointwiseProduct F ⋙ (evaluation D C).obj d ≅ + pointwiseProduct (fun s => F s ⋙ (evaluation _ _).obj d) := + NatIso.ofComponents (fun k => piObjIso _ _) + (fun f => Pi.hom_ext _ _ (by simp [← NatTrans.comp_app])) + +variable [∀ i, HasColimitsOfShape (I i) C] [HasColimitsOfShape (∀ i, I i) C] + +theorem colimitPointwiseProductToProductColimit_app (d : D) : + (colimitPointwiseProductToProductColimit F).app d = + (colimitObjIsoColimitCompEvaluation _ _).hom ≫ + (HasColimit.isoOfNatIso (pointwiseProductCompEvaluation F d)).hom ≫ + colimitPointwiseProductToProductColimit _ ≫ + (Pi.mapIso fun _ => (colimitObjIsoColimitCompEvaluation _ _).symm).hom ≫ + (piObjIso _ _).inv := by + rw [← Iso.inv_comp_eq] + simp only [← Category.assoc] + rw [Iso.eq_comp_inv] + refine Pi.hom_ext _ _ (fun s => colimit.hom_ext (fun k => ?_)) + simp [← NatTrans.comp_app] + +end functorCategory + +section + +variable (C : Type u) [Category.{v} C] + +/-- A category `C` has the `w`-IPC property if the natural morphism +`colim_k (∏ᶜ s ↦ (F s).obj (k s)) ⟶ ∏ᶜ s ↦ colim_k (F s).obj (k s)` is an isomorphism for any +family of functors `F i : I i ⥤ C` with `I i` `w`-small and filtered for all `i`. -/ +class IsIPC [HasProducts.{w} C] [HasFilteredColimitsOfSize.{w} C] : Prop where + /-- `colimitPointwiseProductToProductColimit F` is always an isomorphism. -/ + isIso : ∀ (α : Type w) (I : α → Type w) [∀ i, SmallCategory (I i)] [∀ i, IsFiltered (I i)] + (F : ∀ i, I i ⥤ C), IsIso (colimitPointwiseProductToProductColimit F) + +attribute [instance] IsIPC.isIso + +end + +section types + +variable {α : Type u} {I : α → Type u} [∀ i, SmallCategory (I i)] [∀ i, IsFiltered (I i)] + +theorem Types.isIso_colimitPointwiseProductToProductColimit (F : ∀ i, I i ⥤ Type u) : + IsIso (colimitPointwiseProductToProductColimit F) := by + -- We follow the proof in [Kashiwara2006], Prop. 3.1.11(ii) + refine (isIso_iff_bijective _).2 ⟨fun y y' hy => ?_, fun x => ?_⟩ + · obtain ⟨ky, yk₀, hyk₀⟩ := Types.jointly_surjective' y + obtain ⟨ky', yk₀', hyk₀'⟩ := Types.jointly_surjective' y' + let k := IsFiltered.max ky ky' + let yk : (pointwiseProduct F).obj k := + (pointwiseProduct F).map (IsFiltered.leftToMax ky ky') yk₀ + let yk' : (pointwiseProduct F).obj k := + (pointwiseProduct F).map (IsFiltered.rightToMax ky ky') yk₀' + obtain rfl : y = colimit.ι (pointwiseProduct F) k yk := by + simp only [yk, Types.Colimit.w_apply', hyk₀] + obtain rfl : y' = colimit.ι (pointwiseProduct F) k yk' := by + simp only [yk', Types.Colimit.w_apply', hyk₀'] + dsimp only [pointwiseProduct_obj] at yk yk' + have hch : ∀ (s : α), ∃ (i' : I s) (hi' : k s ⟶ i'), + (F s).map hi' (Pi.π (fun s => (F s).obj (k s)) s yk) = + (F s).map hi' (Pi.π (fun s => (F s).obj (k s)) s yk') := by + intro s + have hy₁ := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) yk + have hy₂ := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) yk' + dsimp only [pointwiseProduct_obj, types_comp_apply] at hy₁ hy₂ + rw [← hy, hy₁, Types.FilteredColimit.colimit_eq_iff] at hy₂ + obtain ⟨i₀, f₀, g₀, h₀⟩ := hy₂ + refine ⟨IsFiltered.coeq f₀ g₀, f₀ ≫ IsFiltered.coeqHom f₀ g₀, ?_⟩ + conv_rhs => rw [IsFiltered.coeq_condition] + simp [h₀] + choose k' f hk' using hch + apply Types.colimit_sound' f f + exact Types.limit_ext' _ _ _ (fun ⟨s⟩ => by simpa using hk' _) + · have hch : ∀ (s : α), ∃ (i : I s) (xi : (F s).obj i), colimit.ι (F s) i xi = + Pi.π (fun s => colimit (F s)) s x := fun s => Types.jointly_surjective' _ + choose k p hk using hch + refine ⟨colimit.ι (pointwiseProduct F) k ((Types.productIso _).inv p), ?_⟩ + refine Types.limit_ext' _ _ _ (fun ⟨s⟩ => ?_) + have := congrFun (ι_colimitPointwiseProductToProductColimit_π F k s) + ((Types.productIso _).inv p) + exact this.trans (by simpa using hk _) + +instance : IsIPC.{u} (Type u) where + isIso _ _ := Types.isIso_colimitPointwiseProductToProductColimit + +end types + +section functorCategory + +variable {C : Type u} [Category.{v} C] + +instance [HasProducts.{w} C] [HasFilteredColimitsOfSize.{w, w} C] [IsIPC.{w} C] {D : Type u₁} + [Category.{v₁} D] : IsIPC.{w} (D ⥤ C) := by + refine ⟨fun β I _ _ F => ?_⟩ + suffices ∀ d, IsIso ((colimitPointwiseProductToProductColimit F).app d) from + NatIso.isIso_of_isIso_app _ + exact fun d => colimitPointwiseProductToProductColimit_app F d ▸ inferInstance + +end functorCategory + +end CategoryTheory.Limits diff --git a/Mathlib/CategoryTheory/Limits/Indization/Products.lean b/Mathlib/CategoryTheory/Limits/Indization/Products.lean new file mode 100644 index 0000000000000..ba9c97fd879ee --- /dev/null +++ b/Mathlib/CategoryTheory/Limits/Indization/Products.lean @@ -0,0 +1,43 @@ +/- +Copyright (c) 2024 Markus Himmel. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Markus Himmel +-/ +import Mathlib.CategoryTheory.Limits.FilteredColimitCommutesProduct +import Mathlib.CategoryTheory.Limits.Indization.FilteredColimits + +/-! +# Ind-objects are closed under products + +We show that if `C` admits products indexed by `α`, then `IsIndObject` is closed under taking +products in `Cᵒᵖ ⥤ Type v` indexed by `α`. This will imply that the functor `Ind C ⥤ Cᵒᵖ ⥤ Type v` +creates products indexed by `α` and that the functor `C ⥤ Ind C` preserves them. + +## References +* [M. Kashiwara, P. Schapira, *Categories and Sheaves*][Kashiwara2006], Prop. 6.1.16(ii) +-/ + +universe v u + +namespace CategoryTheory.Limits + +variable {C : Type u} [Category.{v} C] {α : Type v} + +theorem isIndObject_pi (h : ∀ (g : α → C), IsIndObject (∏ᶜ yoneda.obj ∘ g)) + (f : α → Cᵒᵖ ⥤ Type v) (hf : ∀ a, IsIndObject (f a)) : IsIndObject (∏ᶜ f) := by + let F := fun a => (hf a).presentation.F ⋙ yoneda + suffices (∏ᶜ f ≅ colimit (pointwiseProduct F)) from + (isIndObject_colimit _ _ (fun i => h _)).map this.inv + refine Pi.mapIso (fun s => ?_) ≪≫ (asIso (colimitPointwiseProductToProductColimit F)).symm + exact IsColimit.coconePointUniqueUpToIso (hf s).presentation.isColimit (colimit.isColimit _) + +theorem isIndObject_limit_of_discrete (h : ∀ (g : α → C), IsIndObject (∏ᶜ yoneda.obj ∘ g)) + (F : Discrete α ⥤ Cᵒᵖ ⥤ Type v) (hF : ∀ a, IsIndObject (F.obj a)) : IsIndObject (limit F) := + IsIndObject.map (Pi.isoLimit _).hom (isIndObject_pi h _ (fun a => hF ⟨a⟩)) + +theorem isIndObject_limit_of_discrete_of_hasLimitsOfShape [HasLimitsOfShape (Discrete α) C] + (F : Discrete α ⥤ Cᵒᵖ ⥤ Type v) (hF : ∀ a, IsIndObject (F.obj a)) : IsIndObject (limit F) := + isIndObject_limit_of_discrete (fun g => (isIndObject_limit_comp_yoneda (Discrete.functor g)).map + (HasLimit.isoOfNatIso (Discrete.compNatIsoDiscrete g yoneda)).hom) F hF + +end CategoryTheory.Limits From 3bcc7678efce412065e09ed91f8040020ff052ae Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 10:56:04 +0000 Subject: [PATCH 222/250] chore(zulip_emoji_merge_delegate): exclude #rss channel from search (#19533) --- scripts/zulip_emoji_merge_delegate.py | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index 8fa4e33e74f4c..cbecbdf203c10 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,14 +23,14 @@ site=ZULIP_SITE ) -# Fetch the messages containing the PR number from the public channels. +# Fetch the messages containing the PR number from the public channels (exluding #rss). # There does not seem to be a way to search simultaneously public and private channels. public_response = client.get_messages({ "anchor": "newest", "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channels", "operand": "public"}, + {"operator": "channel", "operand": "rss", "negated": true}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) From 81641565e7805c54b3912cdad6641cd4848266d1 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 21:57:39 +1100 Subject: [PATCH 223/250] bump batteries --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 535ccc0955f60..69fb66fd01c0f 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "3992cab30d93eaaf149f573a0f89c6e158786036", + "rev": "0b7e415ddc73df96a3db8b39aca5fb2306d911e9", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", From ec91d04f5c1a39c8c6434e6535a8bda040b54925 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 22:02:11 +1100 Subject: [PATCH 224/250] bump deps --- lake-manifest.json | 6 +++--- lakefile.lean | 2 +- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/lake-manifest.json b/lake-manifest.json index 69fb66fd01c0f..79489f07f8489 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -35,10 +35,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "1383e72b40dd62a566896a6e348ffe868801b172", + "rev": "5f1a4a085063afd0d0d0fd5dae37d7233a00c0fb", "name": "proofwidgets", "manifestFile": "lake-manifest.json", - "inputRev": "v0.0.46", + "inputRev": "v0.0.47-pre", "inherited": false, "configFile": "lakefile.lean"}, {"url": "https://github.com/leanprover-community/aesop", @@ -65,7 +65,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "0b7e415ddc73df96a3db8b39aca5fb2306d911e9", + "rev": "2076057d018822a2583e6238c157f684708d11ed", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing", diff --git a/lakefile.lean b/lakefile.lean index 60f909ae08789..1faeb0f5bc286 100644 --- a/lakefile.lean +++ b/lakefile.lean @@ -10,7 +10,7 @@ open Lake DSL 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" +require "leanprover-community" / "proofwidgets" @ git "v0.0.47-pre" require "leanprover-community" / "importGraph" @ git "nightly-testing" require "leanprover-community" / "LeanSearchClient" @ git "main" from git "https://github.com/leanprover-community/LeanSearchClient" @ "main" From d98fa7be9fac4bc462268b1ea6926feb25f82ba9 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 22:03:06 +1100 Subject: [PATCH 225/250] deprecation message --- Mathlib/Deprecated/Order.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean index 9f79797c7e46f..93c6a70819f85 100644 --- a/Mathlib/Deprecated/Order.lean +++ b/Mathlib/Deprecated/Order.lean @@ -12,7 +12,7 @@ import Mathlib.Order.Defs.LinearOrder variable {α : Type*} -@[deprecated (since := "2024-11-26")] -- unused in Mathlib +@[deprecated "This was a leftover from Lean 3, and has not been needed." (since := "2024-11-26")] -- unused in Mathlib instance isStrictTotalOrder_of_linearOrder [LinearOrder α] : IsStrictTotalOrder α (· < ·) where irrefl := lt_irrefl trichotomous := lt_trichotomy From d48c7383b7a50f8b38320e6fe3b45bd6c6123b00 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 27 Nov 2024 22:05:32 +1100 Subject: [PATCH 226/250] long line --- Mathlib/Deprecated/Order.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Deprecated/Order.lean b/Mathlib/Deprecated/Order.lean index 93c6a70819f85..65fab91526170 100644 --- a/Mathlib/Deprecated/Order.lean +++ b/Mathlib/Deprecated/Order.lean @@ -12,7 +12,7 @@ import Mathlib.Order.Defs.LinearOrder variable {α : Type*} -@[deprecated "This was a leftover from Lean 3, and has not been needed." (since := "2024-11-26")] -- unused in Mathlib +@[deprecated "This was a leftover from Lean 3, and has not been needed." (since := "2024-11-26")] instance isStrictTotalOrder_of_linearOrder [LinearOrder α] : IsStrictTotalOrder α (· < ·) where irrefl := lt_irrefl trichotomous := lt_trichotomy From 01b31e1d1d58001128368257a32bd3337a5b576e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 11:08:08 +0000 Subject: [PATCH 227/250] chore(Submonoid/Subgroup): rename `subtype_range` to `range_subtype` (#19518) This matches the naming convention. --- Mathlib/Algebra/DirectSum/Basic.lean | 2 +- Mathlib/Algebra/Group/Subgroup/Ker.lean | 21 ++++++++++++------- .../Algebra/Group/Submonoid/Membership.lean | 2 +- .../Algebra/Group/Submonoid/Operations.lean | 14 +++++++++---- Mathlib/GroupTheory/Congruence/Basic.lean | 2 +- Mathlib/GroupTheory/Finiteness.lean | 4 ++-- Mathlib/GroupTheory/Index.lean | 4 ++-- Mathlib/GroupTheory/PGroup.lean | 2 +- Mathlib/GroupTheory/SchurZassenhaus.lean | 6 +++--- .../SpecificGroups/Alternating.lean | 2 +- Mathlib/GroupTheory/Sylow.lean | 6 +++--- 11 files changed, 38 insertions(+), 27 deletions(-) diff --git a/Mathlib/Algebra/DirectSum/Basic.lean b/Mathlib/Algebra/DirectSum/Basic.lean index 50045a6c1ba5b..83ed3a6e295fa 100644 --- a/Mathlib/Algebra/DirectSum/Basic.lean +++ b/Mathlib/Algebra/DirectSum/Basic.lean @@ -375,7 +375,7 @@ def IsInternal {M S : Type*} [DecidableEq ι] [AddCommMonoid M] [SetLike S M] theorem IsInternal.addSubmonoid_iSup_eq_top {M : Type*} [DecidableEq ι] [AddCommMonoid M] (A : ι → AddSubmonoid M) (h : IsInternal A) : iSup A = ⊤ := by - rw [AddSubmonoid.iSup_eq_mrange_dfinsupp_sumAddHom, AddMonoidHom.mrange_eq_top_iff_surjective] + rw [AddSubmonoid.iSup_eq_mrange_dfinsupp_sumAddHom, AddMonoidHom.mrange_eq_top] exact Function.Bijective.surjective h variable {M S : Type*} [AddCommMonoid M] [SetLike S M] [AddSubmonoidClass S M] diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean index 8b9be58076933..5ba75d983b6b7 100644 --- a/Mathlib/Algebra/Group/Subgroup/Ker.lean +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -117,6 +117,9 @@ lemma rangeRestrict_injective_iff {f : G →* N} : Injective f.rangeRestrict ↔ theorem map_range (g : N →* P) (f : G →* N) : f.range.map g = (g.comp f).range := by rw [range_eq_map, range_eq_map]; exact (⊤ : Subgroup G).map_map g f +@[to_additive] +lemma range_comp (g : N →* P) (f : G →* N) : (g.comp f).range = f.range.map g := (map_range ..).symm + @[to_additive] theorem range_eq_top {N} [Group N] {f : G →* N} : f.range = (⊤ : Subgroup N) ↔ Function.Surjective f := @@ -138,10 +141,11 @@ theorem range_one : (1 : G →* N).range = ⊥ := SetLike.ext fun x => by simpa using @comm _ (· = ·) _ 1 x @[to_additive (attr := simp)] -theorem _root_.Subgroup.subtype_range (H : Subgroup G) : H.subtype.range = H := by - rw [range_eq_map, ← SetLike.coe_set_eq, coe_map, Subgroup.coeSubtype] - ext - simp +theorem _root_.Subgroup.range_subtype (H : Subgroup G) : H.subtype.range = H := + SetLike.coe_injective <| (coe_range _).trans <| Subtype.range_coe + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias _root_.Subgroup.subtype_range := Subgroup.range_subtype @[to_additive (attr := simp)] theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : @@ -368,7 +372,7 @@ theorem map_le_range (H : Subgroup G) : map f H ≤ f.range := @[to_additive] theorem map_subtype_le {H : Subgroup G} (K : Subgroup H) : K.map H.subtype ≤ H := - (K.map_le_range H.subtype).trans (le_of_eq H.subtype_range) + (K.map_le_range H.subtype).trans_eq H.range_subtype @[to_additive] theorem ker_le_comap (H : Subgroup N) : f.ker ≤ comap f H := @@ -466,7 +470,7 @@ theorem map_injective_of_ker_le {H K : Subgroup G} (hH : f.ker ≤ H) (hK : f.ke @[to_additive] theorem closure_preimage_eq_top (s : Set G) : closure ((closure s).subtype ⁻¹' s) = ⊤ := by apply map_injective (closure s).subtype_injective - rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, subtype_range, + rw [MonoidHom.map_closure, ← MonoidHom.range_eq_map, range_subtype, Set.image_preimage_eq_of_subset] rw [coeSubtype, Subtype.range_coe_subtype] exact subset_closure @@ -487,7 +491,8 @@ theorem comap_sup_eq (H K : Subgroup N) (hf : Function.Surjective f) : @[to_additive] theorem sup_subgroupOf_eq {H K L : Subgroup G} (hH : H ≤ L) (hK : K ≤ L) : H.subgroupOf L ⊔ K.subgroupOf L = (H ⊔ K).subgroupOf L := - comap_sup_eq_of_le_range L.subtype (hH.trans L.subtype_range.ge) (hK.trans L.subtype_range.ge) + comap_sup_eq_of_le_range L.subtype (hH.trans_eq L.range_subtype.symm) + (hK.trans_eq L.range_subtype.symm) @[to_additive] theorem codisjoint_subgroupOf_sup (H K : Subgroup G) : @@ -501,7 +506,7 @@ theorem subgroupOf_sup (A A' B : Subgroup G) (hA : A ≤ B) (hA' : A' ≤ B) : refine map_injective_of_ker_le B.subtype (ker_le_comap _ _) (le_trans (ker_le_comap B.subtype _) le_sup_left) ?_ - simp only [subgroupOf, map_comap_eq, map_sup, subtype_range] + simp only [subgroupOf, map_comap_eq, map_sup, range_subtype] rw [inf_of_le_right (sup_le hA hA'), inf_of_le_right hA', inf_of_le_right hA] end Subgroup diff --git a/Mathlib/Algebra/Group/Submonoid/Membership.lean b/Mathlib/Algebra/Group/Submonoid/Membership.lean index b48b123c3e191..2eeb9e1f37c14 100644 --- a/Mathlib/Algebra/Group/Submonoid/Membership.lean +++ b/Mathlib/Algebra/Group/Submonoid/Membership.lean @@ -543,7 +543,7 @@ open MonoidHom @[to_additive] theorem sup_eq_range (s t : Submonoid N) : s ⊔ t = mrange (s.subtype.coprod t.subtype) := by rw [mrange_eq_map, ← mrange_inl_sup_mrange_inr, map_sup, map_mrange, coprod_comp_inl, map_mrange, - coprod_comp_inr, range_subtype, range_subtype] + coprod_comp_inr, mrange_subtype, mrange_subtype] @[to_additive] theorem mem_sup {s t : Submonoid N} {x : N} : x ∈ s ⊔ t ↔ ∃ y ∈ s, ∃ z ∈ t, y * z = x := by diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 36ee592c01e9d..86c4066e007df 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -628,6 +628,10 @@ theorem coe_mrange (f : F) : (mrange f : Set N) = Set.range f := theorem mem_mrange {f : F} {y : N} : y ∈ mrange f ↔ ∃ x, f x = y := Iff.rfl +@[to_additive] +lemma mrange_comp {O : Type*} [Monoid O] (f : N →* O) (g : M →* N) : + mrange (f.comp g) = (mrange g).map f := SetLike.coe_injective <| Set.range_comp f _ + @[to_additive] theorem mrange_eq_map (f : F) : mrange f = (⊤ : Submonoid M).map f := Submonoid.copy_eq _ @@ -641,18 +645,18 @@ theorem map_mrange (g : N →* P) (f : M →* N) : f.mrange.map g = mrange (comp simpa only [mrange_eq_map] using (⊤ : Submonoid M).map_map g f @[to_additive] -theorem mrange_eq_top_iff_surjective {f : F} : mrange f = (⊤ : Submonoid N) ↔ Surjective f := +theorem mrange_eq_top {f : F} : mrange f = (⊤ : Submonoid N) ↔ Surjective f := SetLike.ext'_iff.trans <| Iff.trans (by rw [coe_mrange, coe_top]) Set.range_eq_univ @[deprecated (since := "2024-11-11")] -alias mrange_top_iff_surjective := mrange_eq_top_iff_surjective +alias mrange_top_iff_surjective := mrange_eq_top /-- The range of a surjective monoid hom is the whole of the codomain. -/ @[to_additive (attr := simp) "The range of a surjective `AddMonoid` hom is the whole of the codomain."] theorem mrange_eq_top_of_surjective (f : F) (hf : Function.Surjective f) : mrange f = (⊤ : Submonoid N) := - mrange_eq_top_iff_surjective.2 hf + mrange_eq_top.2 hf @[deprecated (since := "2024-11-11")] alias mrange_top_of_surjective := mrange_eq_top_of_surjective @@ -863,9 +867,11 @@ def inclusion {S T : Submonoid M} (h : S ≤ T) : S →* T := S.subtype.codRestrict _ fun x => h x.2 @[to_additive (attr := simp)] -theorem range_subtype (s : Submonoid M) : mrange s.subtype = s := +theorem mrange_subtype (s : Submonoid M) : mrange s.subtype = s := SetLike.coe_injective <| (coe_mrange _).trans <| Subtype.range_coe +@[to_additive (attr := deprecated (since := "2024-11-25"))] alias range_subtype := mrange_subtype + @[to_additive] theorem eq_top_iff' : S = ⊤ ↔ ∀ x : M, x ∈ S := eq_top_iff.trans ⟨fun h m => h <| mem_top m, fun h m _ => h m⟩ diff --git a/Mathlib/GroupTheory/Congruence/Basic.lean b/Mathlib/GroupTheory/Congruence/Basic.lean index 2810c0bd2c101..259eef9b6e583 100644 --- a/Mathlib/GroupTheory/Congruence/Basic.lean +++ b/Mathlib/GroupTheory/Congruence/Basic.lean @@ -173,7 +173,7 @@ variable (x y : M) @[to_additive (attr := simp)] -- Porting note: removed dot notation theorem mrange_mk' : MonoidHom.mrange c.mk' = ⊤ := - MonoidHom.mrange_eq_top_iff_surjective.2 mk'_surjective + MonoidHom.mrange_eq_top.2 mk'_surjective variable {f : M →* P} diff --git a/Mathlib/GroupTheory/Finiteness.lean b/Mathlib/GroupTheory/Finiteness.lean index bd31d6b8743b4..b796773c813c6 100644 --- a/Mathlib/GroupTheory/Finiteness.lean +++ b/Mathlib/GroupTheory/Finiteness.lean @@ -138,7 +138,7 @@ theorem Submonoid.FG.map_injective {M' : Type*} [Monoid M'] {P : Submonoid M} (e @[to_additive (attr := simp)] theorem Monoid.fg_iff_submonoid_fg (N : Submonoid M) : Monoid.FG N ↔ N.FG := by - conv_rhs => rw [← N.range_subtype, MonoidHom.mrange_eq_map] + conv_rhs => rw [← N.mrange_subtype, MonoidHom.mrange_eq_map] exact ⟨fun h => h.out.map N.subtype, fun h => ⟨h.map_injective N.subtype Subtype.coe_injective⟩⟩ @[to_additive] @@ -148,7 +148,7 @@ theorem Monoid.fg_of_surjective {M' : Type*} [Monoid M'] [Monoid.FG M] (f : M obtain ⟨s, hs⟩ := Monoid.fg_def.mp ‹_› use s.image f rwa [Finset.coe_image, ← MonoidHom.map_mclosure, hs, ← MonoidHom.mrange_eq_map, - MonoidHom.mrange_eq_top_iff_surjective] + MonoidHom.mrange_eq_top] @[to_additive] instance Monoid.fg_range {M' : Type*} [Monoid M'] [Monoid.FG M] (f : M →* M') : diff --git a/Mathlib/GroupTheory/Index.lean b/Mathlib/GroupTheory/Index.lean index 0201d5ad30137..95b2a5c3c1451 100644 --- a/Mathlib/GroupTheory/Index.lean +++ b/Mathlib/GroupTheory/Index.lean @@ -80,7 +80,7 @@ theorem index_comap (f : G' →* G) : @[to_additive] theorem relindex_comap (f : G' →* G) (K : Subgroup G') : relindex (comap f H) K = relindex H (map f K) := by - rw [relindex, subgroupOf, comap_comap, index_comap, ← f.map_range, K.subtype_range] + rw [relindex, subgroupOf, comap_comap, index_comap, ← f.map_range, K.range_subtype] variable {H K L} @@ -262,7 +262,7 @@ theorem index_map_of_injective {f : G →* G'} (hf : Function.Injective f) : @[to_additive] theorem index_map_subtype {H : Subgroup G} (K : Subgroup H) : (K.map H.subtype).index = K.index * H.index := by - rw [K.index_map_of_injective H.subtype_injective, H.subtype_range] + rw [K.index_map_of_injective H.subtype_injective, H.range_subtype] @[to_additive] theorem index_eq_card : H.index = Nat.card (G ⧸ H) := diff --git a/Mathlib/GroupTheory/PGroup.lean b/Mathlib/GroupTheory/PGroup.lean index 071291b214ad5..abcab1e3e679d 100644 --- a/Mathlib/GroupTheory/PGroup.lean +++ b/Mathlib/GroupTheory/PGroup.lean @@ -240,7 +240,7 @@ theorem to_inf_right {H K : Subgroup G} (hK : IsPGroup p K) : IsPGroup p (H ⊓ theorem map {H : Subgroup G} (hH : IsPGroup p H) {K : Type*} [Group K] (ϕ : G →* K) : IsPGroup p (H.map ϕ) := by - rw [← H.subtype_range, MonoidHom.map_range] + rw [← H.range_subtype, MonoidHom.map_range] exact hH.of_surjective (ϕ.restrict H).rangeRestrict (ϕ.restrict H).rangeRestrict_surjective theorem comap_of_ker_isPGroup {H : Subgroup G} (hH : IsPGroup p H) {K : Type*} [Group K] diff --git a/Mathlib/GroupTheory/SchurZassenhaus.lean b/Mathlib/GroupTheory/SchurZassenhaus.lean index acdf40f97cdde..e39e21cc46912 100644 --- a/Mathlib/GroupTheory/SchurZassenhaus.lean +++ b/Mathlib/GroupTheory/SchurZassenhaus.lean @@ -174,7 +174,7 @@ private theorem step1 (K : Subgroup G) (hK : K ⊔ N = ⊤) : K = ⊤ := by replace hH : Nat.card (H.map K.subtype) = N.index := by rw [← relindex_bot_left, ← relindex_comap, MonoidHom.comap_bot, Subgroup.ker_subtype, relindex_bot_left, ← IsComplement'.index_eq_card (IsComplement'.symm hH), index_comap, - subtype_range, ← relindex_sup_right, hK, relindex_top_right] + range_subtype, ← relindex_sup_right, hK, relindex_top_right] have h7 : Nat.card N * Nat.card (H.map K.subtype) = Nat.card G := by rw [hH, ← N.index_mul_card, mul_comm] have h8 : (Nat.card N).Coprime (Nat.card (H.map K.subtype)) := by @@ -220,7 +220,7 @@ private theorem step3 (K : Subgroup N) [(K.map N.subtype).Normal] : K = ⊥ ∨ conv at key => rhs rhs - rw [← N.subtype_range, N.subtype.range_eq_map] + rw [← N.range_subtype, N.subtype.range_eq_map] have inj := map_injective N.subtype_injective rwa [inj.eq_iff, inj.eq_iff] at key @@ -239,7 +239,7 @@ include h2 in private theorem step6 : IsPGroup (Nat.card N).minFac N := by haveI : Fact (Nat.card N).minFac.Prime := ⟨step4 h1 h3⟩ refine Sylow.nonempty.elim fun P => P.2.of_surjective P.1.subtype ?_ - rw [← MonoidHom.range_eq_top, subtype_range] + rw [← MonoidHom.range_eq_top, range_subtype] haveI : (P.1.map N.subtype).Normal := normalizer_eq_top.mp (step1 h1 h2 h3 (P.1.map N.subtype).normalizer P.normalizer_sup_eq_top) exact (step3 h1 h2 h3 P.1).resolve_left (step5 h1 h3) diff --git a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean index 8b0f82d484a0d..7dd9feaf90d4d 100644 --- a/Mathlib/GroupTheory/SpecificGroups/Alternating.lean +++ b/Mathlib/GroupTheory/SpecificGroups/Alternating.lean @@ -168,7 +168,7 @@ theorem IsThreeCycle.alternating_normalClosure (h5 : 5 ≤ Fintype.card α) {f : (by have hi : Function.Injective (alternatingGroup α).subtype := Subtype.coe_injective refine eq_top_iff.1 (map_injective hi (le_antisymm (map_mono le_top) ?_)) - rw [← MonoidHom.range_eq_map, subtype_range, normalClosure, MonoidHom.map_closure] + rw [← MonoidHom.range_eq_map, range_subtype, normalClosure, MonoidHom.map_closure] refine (le_of_eq closure_three_cycles_eq_alternating.symm).trans (closure_mono ?_) intro g h obtain ⟨c, rfl⟩ := isConj_iff.1 (isConj_iff_cycleType_eq.2 (hf.trans h.symm)) diff --git a/Mathlib/GroupTheory/Sylow.lean b/Mathlib/GroupTheory/Sylow.lean index fdfa524a631b9..6b4655caeb3b5 100644 --- a/Mathlib/GroupTheory/Sylow.lean +++ b/Mathlib/GroupTheory/Sylow.lean @@ -140,7 +140,7 @@ theorem coe_comapOfInjective (hϕ : Function.Injective ϕ) (h : P ≤ ϕ.range) /-- A sylow subgroup of G is also a sylow subgroup of a subgroup of G. -/ protected def subtype (h : P ≤ N) : Sylow p N := - P.comapOfInjective N.subtype Subtype.coe_injective (by rwa [subtype_range]) + P.comapOfInjective N.subtype Subtype.coe_injective (by rwa [range_subtype]) @[simp] theorem coe_subtype (h : P ≤ N) : P.subtype h = subgroupOf P N := @@ -756,9 +756,9 @@ theorem normalizer_normalizer {p : ℕ} [Fact p.Prime] [Finite (Sylow p G)] (P : have := normal_of_normalizer_normal (P.subtype (le_normalizer.trans le_normalizer)) simp_rw [← normalizer_eq_top, coe_subtype, ← subgroupOf_normalizer_eq le_normalizer, ← subgroupOf_normalizer_eq le_rfl, subgroupOf_self] at this - rw [← subtype_range P.normalizer.normalizer, MonoidHom.range_eq_map, + rw [← range_subtype P.normalizer.normalizer, MonoidHom.range_eq_map, ← this trivial] - exact map_comap_eq_self (le_normalizer.trans (ge_of_eq (subtype_range _))) + exact map_comap_eq_self (le_normalizer.trans (ge_of_eq (range_subtype _))) theorem normal_of_all_max_subgroups_normal [Finite G] (hnc : ∀ H : Subgroup G, IsCoatom H → H.Normal) {p : ℕ} [Fact p.Prime] [Finite (Sylow p G)] From cb86c7503d19e5d7a123067aec02599a8eaad66d Mon Sep 17 00:00:00 2001 From: Christian Merten Date: Wed, 27 Nov 2024 11:40:18 +0000 Subject: [PATCH 228/250] feat(CategoryTheory): pushforward pullback adjunction for `P.Over Q X` (#19271) --- Mathlib.lean | 1 + .../Limits/MorphismProperty.lean | 3 +- .../MorphismProperty/Basic.lean | 4 + .../MorphismProperty/Comma.lean | 101 ++++++++++++- .../MorphismProperty/OverAdjunction.lean | 133 ++++++++++++++++++ 5 files changed, 238 insertions(+), 4 deletions(-) create mode 100644 Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3e568ce173a2..a843c7987f61d 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1991,6 +1991,7 @@ import Mathlib.CategoryTheory.MorphismProperty.Concrete import Mathlib.CategoryTheory.MorphismProperty.Factorization import Mathlib.CategoryTheory.MorphismProperty.IsInvertedBy import Mathlib.CategoryTheory.MorphismProperty.Limits +import Mathlib.CategoryTheory.MorphismProperty.OverAdjunction import Mathlib.CategoryTheory.MorphismProperty.Representable import Mathlib.CategoryTheory.NatIso import Mathlib.CategoryTheory.NatTrans diff --git a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean index b4bb842c72dd2..48a845bacb238 100644 --- a/Mathlib/CategoryTheory/Limits/MorphismProperty.lean +++ b/Mathlib/CategoryTheory/Limits/MorphismProperty.lean @@ -119,9 +119,8 @@ instance [P.ContainsIdentities] (Y : P.Over ⊤ X) : uniq a := by ext · simp only [mk_left, Hom.hom_left, homMk_hom, Over.homMk_left] - rw [← Over.w a.hom] + rw [← Over.w a] simp only [mk_left, Functor.const_obj_obj, Hom.hom_left, mk_hom, Category.comp_id] - · rfl /-- `X ⟶ X` is the terminal object of `P.Over ⊤ X`. -/ def mkIdTerminal [P.ContainsIdentities] : diff --git a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean index cf9f01545d900..e8a96be14d9ab 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Basic.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Basic.lean @@ -157,6 +157,10 @@ lemma RespectsIso.precomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} ( [IsIso e] (f : Y ⟶ Z) (hf : P f) : P (e ≫ f) := RespectsLeft.precomp (Q := isomorphisms C) e ‹IsIso e› f hf +instance : RespectsIso (⊤ : MorphismProperty C) where + precomp _ _ _ _ := trivial + postcomp _ _ _ _ := trivial + lemma RespectsIso.postcomp (P : MorphismProperty C) [P.RespectsIso] {X Y Z : C} (e : Y ⟶ Z) [IsIso e] (f : X ⟶ Y) (hf : P f) : P (f ≫ e) := RespectsRight.postcomp (Q := isomorphisms C) e ‹IsIso e› f hf diff --git a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean index 6b093565af86f..3dd55847c9890 100644 --- a/Mathlib/CategoryTheory/MorphismProperty/Comma.lean +++ b/Mathlib/CategoryTheory/MorphismProperty/Comma.lean @@ -79,10 +79,8 @@ lemma Hom.hom_mk {X Y : P.Comma L R Q W} (f : CommaMorphism X.toComma Y.toComma) (hf) (hg) : Comma.Hom.hom ⟨f, hf, hg⟩ = f := rfl -@[simp] lemma Hom.hom_left {X Y : P.Comma L R Q W} (f : Comma.Hom X Y) : f.hom.left = f.left := rfl -@[simp] lemma Hom.hom_right {X Y : P.Comma L R Q W} (f : Comma.Hom X Y) : f.hom.right = f.right := rfl /-- See Note [custom simps projection] -/ @@ -133,6 +131,14 @@ lemma id_hom (X : P.Comma L R Q W) : (𝟙 X : X ⟶ X).hom = 𝟙 X.toComma := lemma comp_hom {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : (f ≫ g).hom = f.hom ≫ g.hom := rfl +@[reassoc] +lemma comp_left {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).left = f.left ≫ g.left := rfl + +@[reassoc] +lemma comp_right {X Y Z : P.Comma L R Q W} (f : X ⟶ Y) (g : Y ⟶ Z) : + (f ≫ g).right = f.right ≫ g.right := rfl + /-- If `i` is an isomorphism in `Comma L R`, it is also a morphism in `P.Comma L R Q W`. -/ @[simps hom] def homFromCommaOfIsIso [Q.RespectsIso] [W.RespectsIso] {X Y : P.Comma L R Q W} @@ -205,6 +211,63 @@ def forgetFullyFaithful : (forget L R P ⊤ ⊤).FullyFaithful where instance : (forget L R P ⊤ ⊤).Full := Functor.FullyFaithful.full (forgetFullyFaithful L R P) +section + +variable {L R} + +@[simp] +lemma eqToHom_left {X Y : P.Comma L R Q W} (h : X = Y) : + (eqToHom h).left = eqToHom (by rw [h]) := by + subst h + rfl + +@[simp] +lemma eqToHom_right {X Y : P.Comma L R Q W} (h : X = Y) : + (eqToHom h).right = eqToHom (by rw [h]) := by + subst h + rfl + +end + +section Functoriality + +variable {L R P Q W} +variable {L₁ L₂ L₃ : A ⥤ T} {R₁ R₂ R₃ : B ⥤ T} + +/-- Lift a functor `F : C ⥤ Comma L R` to the subcategory `P.Comma L R Q W` under +suitable assumptions on `F`. -/ +@[simps obj_toComma map_hom] +def lift {C : Type*} [Category C] (F : C ⥤ Comma L R) + (hP : ∀ X, P (F.obj X).hom) + (hQ : ∀ {X Y} (f : X ⟶ Y), Q (F.map f).left) + (hW : ∀ {X Y} (f : X ⟶ Y), W (F.map f).right) : + C ⥤ P.Comma L R Q W where + obj X := + { __ := F.obj X + prop := hP X } + map {X Y} f := + { __ := F.map f + prop_hom_left := hQ f + prop_hom_right := hW f } + +variable (R) in +/-- A natural transformation `L₁ ⟶ L₂` induces a functor `P.Comma L₂ R Q W ⥤ P.Comma L₁ R Q W`. -/ +@[simps!] +def mapLeft (l : L₁ ⟶ L₂) (hl : ∀ X : P.Comma L₂ R Q W, P (l.app X.left ≫ X.hom)) : + P.Comma L₂ R Q W ⥤ P.Comma L₁ R Q W := + lift (forget _ _ _ _ _ ⋙ CategoryTheory.Comma.mapLeft R l) hl + (fun f ↦ f.prop_hom_left) (fun f ↦ f.prop_hom_right) + +variable (L) in +/-- A natural transformation `R₁ ⟶ R₂` induces a functor `P.Comma L R₁ Q W ⥤ P.Comma L R₂ Q W`. -/ +@[simps!] +def mapRight (r : R₁ ⟶ R₂) (hr : ∀ X : P.Comma L R₁ Q W, P (X.hom ≫ r.app X.right)) : + P.Comma L R₁ Q W ⥤ P.Comma L R₂ Q W := + lift (forget _ _ _ _ _ ⋙ CategoryTheory.Comma.mapRight L r) hr + (fun f ↦ f.prop_hom_left) (fun f ↦ f.prop_hom_right) + +end Functoriality + end Comma end Comma @@ -251,6 +314,23 @@ protected def Over.homMk {A B : P.Over Q X} (f : A.left ⟶ B.left) prop_hom_left := hf prop_hom_right := trivial +/-- Make an isomorphism in `P.Over Q X` from an isomorphism in `T` with compatibilities. -/ +@[simps! hom_left inv_left] +protected def Over.isoMk [Q.RespectsIso] {A B : P.Over Q X} (f : A.left ≅ B.left) + (w : f.hom ≫ B.hom = A.hom := by aesop_cat) : A ≅ B := + Comma.isoMk f (Discrete.eqToIso' rfl) + +@[ext] +lemma Over.Hom.ext {A B : P.Over Q X} {f g : A ⟶ B} (h : f.left = g.left) : f = g := by + ext + · exact h + · simp + +@[reassoc] +lemma Over.w {A B : P.Over Q X} (f : A ⟶ B) : + f.left ≫ B.hom = A.hom := by + simp + end Over section Under @@ -295,6 +375,23 @@ protected def Under.homMk {A B : P.Under Q X} (f : A.right ⟶ B.right) prop_hom_left := trivial prop_hom_right := hf +/-- Make an isomorphism in `P.Under Q X` from an isomorphism in `T` with compatibilities. -/ +@[simps! hom_right inv_right] +protected def Under.isoMk [Q.RespectsIso] {A B : P.Under Q X} (f : A.right ≅ B.right) + (w : A.hom ≫ f.hom = B.hom := by aesop_cat) : A ≅ B := + Comma.isoMk (Discrete.eqToIso' rfl) f + +@[ext] +lemma Under.Hom.ext {A B : P.Under Q X} {f g : A ⟶ B} (h : f.right = g.right) : f = g := by + ext + · simp + · exact h + +@[reassoc] +lemma Under.w {A B : P.Under Q X} (f : A ⟶ B) : + A.hom ≫ f.right = B.hom := by + simp + end Under end CategoryTheory.MorphismProperty diff --git a/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean b/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean new file mode 100644 index 0000000000000..4a41ebda2533c --- /dev/null +++ b/Mathlib/CategoryTheory/MorphismProperty/OverAdjunction.lean @@ -0,0 +1,133 @@ +/- +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.CategoryTheory.MorphismProperty.Comma +import Mathlib.CategoryTheory.Adjunction.Over +import Mathlib.CategoryTheory.MorphismProperty.Limits + +/-! +# Adjunction of pushforward and pullback in `P.Over Q X` + +A morphism `f : X ⟶ Y` defines two functors: + +- `Over.map`: post-composition with `f` +- `Over.pullback`: base-change along `f` + +These are adjoint under suitable assumptions on `P` and `Q`. + +-/ + +namespace CategoryTheory.MorphismProperty + +open Limits + +variable {T : Type*} [Category T] (P Q : MorphismProperty T) [Q.IsMultiplicative] +variable {X Y : T} (f : X ⟶ Y) + +section Map + +variable {P} [P.IsStableUnderComposition] (hPf : P f) + +variable {f} + +/-- If `P` is stable under composition and `f : X ⟶ Y` satisfies `P`, +this is the functor `P.Over Q X ⥤ P.Over Q Y` given by composing with `f`. -/ +@[simps! obj_left obj_hom map_left] +def Over.map : P.Over Q X ⥤ P.Over Q Y := + Comma.mapRight _ (Discrete.natTrans fun _ ↦ f) <| fun X ↦ P.comp_mem _ _ X.prop hPf + +lemma Over.map_comp {X Y Z : T} {f : X ⟶ Y} (hf : P f) {g : Y ⟶ Z} (hg : P g) : + map Q (P.comp_mem f g hf hg) = map Q hf ⋙ map Q hg := by + fapply Functor.ext + · simp [map, Comma.mapRight, CategoryTheory.Comma.mapRight, Comma.lift] + · intro U V k + ext + simp + +/-- `Over.map` commutes with composition. -/ +@[simps! hom_app_left inv_app_left] +def Over.mapComp {X Y Z : T} {f : X ⟶ Y} (hf : P f) {g : Y ⟶ Z} (hg : P g) [Q.RespectsIso] : + map Q (P.comp_mem f g hf hg) ≅ map Q hf ⋙ map Q hg := + NatIso.ofComponents (fun X ↦ Over.isoMk (Iso.refl _)) + +end Map + +section Pullback + +variable [HasPullbacks T] [P.IsStableUnderBaseChange] [Q.IsStableUnderBaseChange] + +/-- If `P` and `Q` are stable under base change and pullbacks exist in `T`, +this is the functor `P.Over Q Y ⥤ P.Over Q X` given by base change along `f`. -/ +@[simps! obj_left obj_hom map_left] +noncomputable def Over.pullback : P.Over Q Y ⥤ P.Over Q X where + obj A := + { __ := (CategoryTheory.Over.pullback f).obj A.toComma + prop := P.pullback_snd _ _ A.prop } + map {A B} g := + { __ := (CategoryTheory.Over.pullback f).map g.toCommaMorphism + prop_hom_left := Q.baseChange_map f g.toCommaMorphism g.prop_hom_left + prop_hom_right := trivial } + +variable {P} {Q} + +/-- `Over.pullback` commutes with composition. -/ +@[simps! hom_app_left inv_app_left] +noncomputable def Over.pullbackComp [Q.RespectsIso] {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) : + Over.pullback P Q (f ≫ g) ≅ Over.pullback P Q g ⋙ Over.pullback P Q f := + NatIso.ofComponents + (fun X ↦ Over.isoMk ((pullbackLeftPullbackSndIso X.hom g f).symm) (by simp)) + +lemma Over.pullbackComp_left_fst_fst [Q.RespectsIso] {X Y Z : T} (f : X ⟶ Y) (g : Y ⟶ Z) + (A : P.Over Q Z) : + ((Over.pullbackComp f g).hom.app A).left ≫ + pullback.fst (pullback.snd A.hom g) f ≫ pullback.fst A.hom g = + pullback.fst A.hom (f ≫ g) := by + simp + +/-- If `f = g`, then base change along `f` is naturally isomorphic to base change along `g`. -/ +noncomputable def Over.pullbackCongr {X Y : T} {f g : X ⟶ Y} (h : f = g) : + Over.pullback P Q f ≅ Over.pullback P Q g := + NatIso.ofComponents (fun X ↦ eqToIso (by rw [h])) + +@[reassoc (attr := simp)] +lemma Over.pullbackCongr_hom_app_left_fst {X Y : T} {f g : X ⟶ Y} (h : f = g) (A : P.Over Q Y) : + ((Over.pullbackCongr h).hom.app A).left ≫ pullback.fst A.hom g = + pullback.fst A.hom f := by + subst h + simp [pullbackCongr] + +end Pullback + +section Adjunction + +variable [P.IsStableUnderComposition] [P.IsStableUnderBaseChange] + [Q.IsStableUnderBaseChange] [HasPullbacks T] + +/-- `P.Over.map` is left adjoint to `P.Over.pullback` if `f` satisfies `P`. -/ +noncomputable def Over.mapPullbackAdj [Q.HasOfPostcompProperty Q] (hPf : P f) (hQf : Q f) : + Over.map Q hPf ⊣ Over.pullback P Q f := + Adjunction.mkOfHomEquiv + { homEquiv := fun A B ↦ + { toFun := fun g ↦ + Over.homMk (pullback.lift g.left A.hom <| by simp) (by simp) <| by + apply Q.of_postcomp (W' := Q) + · exact Q.pullback_fst B.hom f hQf + · simpa using g.prop_hom_left + invFun := fun h ↦ Over.homMk (h.left ≫ pullback.fst B.hom f) + (by + simp only [map_obj_left, Functor.const_obj_obj, pullback_obj_left, Functor.id_obj, + Category.assoc, pullback.condition, map_obj_hom, ← pullback_obj_hom, Over.w_assoc]) + (Q.comp_mem _ _ h.prop_hom_left (Q.pullback_fst _ _ hQf)) + left_inv := by aesop_cat + right_inv := fun h ↦ by + ext + dsimp + ext + · simp + · simpa using h.w.symm } } + +end Adjunction + +end CategoryTheory.MorphismProperty From 6b724bf9d816ee5143e5b8b8e8be69b40b820656 Mon Sep 17 00:00:00 2001 From: blizzard_inc Date: Wed, 27 Nov 2024 11:40:20 +0000 Subject: [PATCH 229/250] fix(scripts/zulip_emoji_merge_delegate.py): fix the emoji reaction script (#19535) change the script to use `True` rather than `true`, fixing an error. --- scripts/zulip_emoji_merge_delegate.py | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index cbecbdf203c10..91836a41e49ff 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -30,7 +30,7 @@ "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channel", "operand": "rss", "negated": true}, + {"operator": "channel", "operand": "rss", "negated": True}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) From 671c09054c3b746a0d22ca24711dac1068eb282a Mon Sep 17 00:00:00 2001 From: grunweg Date: Wed, 27 Nov 2024 12:13:29 +0000 Subject: [PATCH 230/250] chore: move style linters to `Style` (#19529) This makes for a more coherent split of the linter files. Third attempt after #15942 and #16510. --- Mathlib/Tactic/Linter/Lint.lean | 337 +---------------------------- Mathlib/Tactic/Linter/Style.lean | 355 ++++++++++++++++++++++++++++++- MathlibTest/Lint.lean | 221 ------------------- MathlibTest/LintStyle.lean | 218 ++++++++++++++++++- MathlibTest/LongFile.lean | 2 +- 5 files changed, 573 insertions(+), 560 deletions(-) diff --git a/Mathlib/Tactic/Linter/Lint.lean b/Mathlib/Tactic/Linter/Lint.lean index 280dad9777e59..67ccae0115ab1 100644 --- a/Mathlib/Tactic/Linter/Lint.lean +++ b/Mathlib/Tactic/Linter/Lint.lean @@ -9,7 +9,8 @@ import Mathlib.Tactic.DeclarationNames /-! # Linters for Mathlib -In this file we define additional linters for mathlib. +In this file we define additional linters for mathlib, +which concern the *behaviour* of the linted code, and not issues of code style or formatting. Perhaps these should be moved to Batteries in the future. -/ @@ -95,336 +96,4 @@ def dupNamespace : Linter where run := withSetOptionIn fun stx ↦ do initialize addLinter dupNamespace -end DupNamespaceLinter - -/-! -# The "missing end" linter - -The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. -It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). --/ - -open Lean Elab Command - -/-- The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. -It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). --/ -register_option linter.style.missingEnd : Bool := { - defValue := false - descr := "enable the missing end linter" -} - -namespace Style.missingEnd - -@[inherit_doc Mathlib.Linter.linter.style.missingEnd] -def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do - -- Only run this linter at the end of a module. - unless stx.isOfKind ``Lean.Parser.Command.eoi do return - if Linter.getLinterValue linter.style.missingEnd (← getOptions) && - !(← MonadState.get).messages.hasErrors then - let sc ← getScopes - -- The last scope is always the "base scope", corresponding to no active `section`s or - -- `namespace`s. We are interested in any *other* unclosed scopes. - if sc.length == 1 then return - let ends := sc.dropLast.map fun s ↦ (s.header, s.isNoncomputable) - -- If the outermost scope corresponds to a `noncomputable section`, we ignore it. - let ends := if ends.getLast!.2 then ends.dropLast else ends - -- If there are any further un-closed scopes, we emit a warning. - if !ends.isEmpty then - let ending := (ends.map Prod.fst).foldl (init := "") fun a b ↦ - a ++ s!"\n\nend{if b == "" then "" else " "}{b}" - Linter.logLint linter.style.missingEnd stx - m!"unclosed sections or namespaces; expected: '{ending}'" - -initialize addLinter missingEndLinter - -end Style.missingEnd - -/-! -# The `cdot` linter - -The `cdot` linter is a syntax-linter that flags uses of the "cdot" `·` that are achieved -by typing a character different from `·`. -For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. -It also flags "isolated cdots", i.e. when the `·` is on its own line. --/ - -/-- -The `cdot` linter flags uses of the "cdot" `·` that are achieved by typing a character -different from `·`. -For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. -It also flags "isolated cdots", i.e. when the `·` is on its own line. -/ -register_option linter.style.cdot : Bool := { - defValue := false - descr := "enable the `cdot` linter" -} - -/-- `isCDot? stx` checks whether `stx` is a `Syntax` node corresponding to a `cdot` typed with -the character `·`. -/ -def isCDot? : Syntax → Bool - | .node _ ``cdotTk #[.node _ `patternIgnore #[.node _ _ #[.atom _ v]]] => v == "·" - | .node _ ``Lean.Parser.Term.cdot #[.atom _ v] => v == "·" - | _ => false - -/-- -`findCDot stx` extracts from `stx` the syntax nodes of `kind` `Lean.Parser.Term.cdot` or `cdotTk`. -/ -partial -def findCDot : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findCDot).flatten - match kind with - | ``Lean.Parser.Term.cdot | ``cdotTk => dargs.push stx - | _ => dargs - |_ => #[] - -/-- `unwanted_cdot stx` returns an array of syntax atoms within `stx` -corresponding to `cdot`s that are not written with the character `·`. -This is precisely what the `cdot` linter flags. --/ -def unwanted_cdot (stx : Syntax) : Array Syntax := - (findCDot stx).filter (!isCDot? ·) - -namespace Style - -@[inherit_doc linter.style.cdot] -def cdotLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.cdot (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in unwanted_cdot stx do - Linter.logLint linter.style.cdot s - m!"Please, use '·' (typed as `\\.`) instead of '{s}' as 'cdot'." - -- We also check for isolated cdot's, i.e. when the cdot is on its own line. - for cdot in Mathlib.Linter.findCDot stx do - match cdot.find? (·.isOfKind `token.«· ») with - | some (.node _ _ #[.atom (.original _ _ afterCDot _) _]) => - if (afterCDot.takeWhile (·.isWhitespace)).contains '\n' then - logWarningAt cdot <| .tagged linter.style.cdot.name - m!"This central dot `·` is isolated; please merge it with the next line." - | _ => return - -initialize addLinter cdotLinter - -end Style - -/-! -# The `dollarSyntax` linter - -The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. -These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. --/ - -/-- The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. -These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. -/ -register_option linter.style.dollarSyntax : Bool := { - defValue := false - descr := "enable the `dollarSyntax` linter" -} - -namespace Style.dollarSyntax - -/-- `findDollarSyntax stx` extracts from `stx` the syntax nodes of `kind` `$`. -/ -partial -def findDollarSyntax : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findDollarSyntax).flatten - match kind with - | ``«term_$__» => dargs.push stx - | _ => dargs - |_ => #[] - -@[inherit_doc linter.style.dollarSyntax] -def dollarSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.dollarSyntax (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in findDollarSyntax stx do - Linter.logLint linter.style.dollarSyntax s - m!"Please use '<|' instead of '$' for the pipe operator." - -initialize addLinter dollarSyntaxLinter - -end Style.dollarSyntax - -/-! -# The `lambdaSyntax` linter - -The `lambdaSyntax` linter is a syntax linter that flags uses of the symbol `λ` to define anonymous -functions, as opposed to the `fun` keyword. These are syntactically equivalent; mathlib style -prefers the latter as it is considered more readable. --/ - -/-- -The `lambdaSyntax` linter flags uses of the symbol `λ` to define anonymous functions. -This is syntactically equivalent to the `fun` keyword; mathlib style prefers using the latter. --/ -register_option linter.style.lambdaSyntax : Bool := { - defValue := false - descr := "enable the `lambdaSyntax` linter" -} - -namespace Style.lambdaSyntax - -/-- -`findLambdaSyntax stx` extracts from `stx` all syntax nodes of `kind` `Term.fun`. -/ -partial -def findLambdaSyntax : Syntax → Array Syntax - | stx@(.node _ kind args) => - let dargs := (args.map findLambdaSyntax).flatten - match kind with - | ``Parser.Term.fun => dargs.push stx - | _ => dargs - |_ => #[] - -@[inherit_doc linter.style.lambdaSyntax] -def lambdaSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.lambdaSyntax (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - for s in findLambdaSyntax stx do - if let .atom _ "λ" := s[0] then - Linter.logLint linter.style.lambdaSyntax s[0] m!"\ - Please use 'fun' and not 'λ' to define anonymous functions.\n\ - The 'λ' syntax is deprecated in mathlib4." - -initialize addLinter lambdaSyntaxLinter - -end Style.lambdaSyntax - -/-! -# The "longFile" linter - -The "longFile" linter emits a warning on files which are longer than a certain number of lines -(1500 by default). --/ - -/-- -The "longFile" linter emits a warning on files which are longer than a certain number of lines -(`linter.style.longFileDefValue` by default on mathlib, no limit for downstream projects). -If this option is set to `N` lines, the linter warns once a file has more than `N` lines. -A value of `0` silences the linter entirely. --/ -register_option linter.style.longFile : Nat := { - defValue := 0 - descr := "enable the longFile linter" -} - -/-- The number of lines that the `longFile` linter considers the default. -/ -register_option linter.style.longFileDefValue : Nat := { - defValue := 1500 - descr := "a soft upper bound on the number of lines of each file" -} - -namespace Style.longFile - -@[inherit_doc Mathlib.Linter.linter.style.longFile] -def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do - let linterBound := linter.style.longFile.get (← getOptions) - if linterBound == 0 then - return - let defValue := linter.style.longFileDefValue.get (← getOptions) - let smallOption := match stx with - | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue - | _ => false - if smallOption then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"The default value of the `longFile` linter is {defValue}.\n\ - The current value of {linterBound} does not exceed the allowed bound.\n\ - Please, remove the `set_option linter.style.longFile {linterBound}`." - else - -- Thanks to the above check, the linter option is either not set (and hence equal - -- to the default) or set to some value *larger* than the default. - -- `Parser.isTerminalCommand` allows `stx` to be `#exit`: this is useful for tests. - unless Parser.isTerminalCommand stx do return - -- We exclude `Mathlib.lean` from the linter: it exceeds linter's default number of allowed - -- lines, and it is an auto-generated import-only file. - -- TODO: if there are more such files, revise the implementation. - if (← getMainModule) == `Mathlib then return - if let some init := stx.getTailPos? then - -- the last line: we subtract 1, since the last line is expected to be empty - let lastLine := ((← getFileMap).toPosition init).line - -- In this case, the file has an allowed length, and the linter option is unnecessarily set. - if lastLine ≤ defValue && defValue < linterBound then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"The default value of the `longFile` linter is {defValue}.\n\ - This file is {lastLine} lines long which does not exceed the allowed bound.\n\ - Please, remove the `set_option linter.style.longFile {linterBound}`." - else - -- `candidate` is divisible by `100` and satisfies `lastLine + 100 < candidate ≤ lastLine + 200` - -- note that either `lastLine ≤ defValue` and `defValue = linterBound` hold or - -- `candidate` is necessarily bigger than `lastLine` and hence bigger than `defValue` - let candidate := (lastLine / 100) * 100 + 200 - let candidate := max candidate defValue - -- In this case, the file is longer than the default and also than what the option says. - if defValue ≤ linterBound && linterBound < lastLine then - logWarningAt stx <| .tagged linter.style.longFile.name - m!"This file is {lastLine} lines long, but the limit is {linterBound}.\n\n\ - You can extend the allowed length of the file using \ - `set_option linter.style.longFile {candidate}`.\n\ - You can completely disable this linter by setting the length limit to `0`." - else - -- Finally, the file exceeds the default value, but not the option: we only allow the value - -- of the option to be `candidate` or `candidate + 100`. - -- In particular, this flags any option that is set to an unnecessarily high value. - if linterBound == candidate || linterBound + 100 == candidate then return - else - logWarningAt stx <| .tagged linter.style.longFile.name - m!"This file is {lastLine} lines long. \ - The current limit is {linterBound}, but it is expected to be {candidate}:\n\ - `set_option linter.style.longFile {candidate}`." - -initialize addLinter longFileLinter - -end Style.longFile - -/-! # The "longLine linter" -/ - -/-- The "longLine" linter emits a warning on lines longer than 100 characters. -We allow lines containing URLs to be longer, though. -/ -register_option linter.style.longLine : Bool := { - defValue := false - descr := "enable the longLine linter" -} - -namespace Style.longLine - -@[inherit_doc Mathlib.Linter.linter.style.longLine] -def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do - unless Linter.getLinterValue linter.style.longLine (← getOptions) do - return - if (← MonadState.get).messages.hasErrors then - return - -- The linter ignores the `#guard_msgs` command, in particular its doc-string. - -- The linter still lints the message guarded by `#guard_msgs`. - if stx.isOfKind ``Lean.guardMsgsCmd then - return - -- if the linter reached the end of the file, then we scan the `import` syntax instead - let stx := ← do - if stx.isOfKind ``Lean.Parser.Command.eoi then - let fileMap ← getFileMap - -- `impMods` is the syntax for the modules imported in the current file - let (impMods, _) ← Parser.parseHeader - { input := fileMap.source, fileName := ← getFileName, fileMap := fileMap } - return impMods - else return stx - let sstr := stx.getSubstring? - let fm ← getFileMap - let longLines := ((sstr.getD default).splitOn "\n").filter fun line ↦ - (100 < (fm.toPosition line.stopPos).column) - for line in longLines do - if (line.splitOn "http").length ≤ 1 then - let stringMsg := if line.contains '"' then - "\nYou can use \"string gaps\" to format long strings: within a string quotation, \ - using a '\\' at the end of a line allows you to continue the string on the following \ - line, removing all intervening whitespace." - else "" - Linter.logLint linter.style.longLine (.ofRange ⟨line.startPos, line.stopPos⟩) - m!"This line exceeds the 100 character limit, please shorten it!{stringMsg}" -initialize addLinter longLineLinter - -end Style.longLine - -end Mathlib.Linter +end Mathlib.Linter.DupNamespaceLinter diff --git a/Mathlib/Tactic/Linter/Style.lean b/Mathlib/Tactic/Linter/Style.lean index 112d95c09ea65..65272ba2d6d0a 100644 --- a/Mathlib/Tactic/Linter/Style.lean +++ b/Mathlib/Tactic/Linter/Style.lean @@ -12,10 +12,29 @@ import Mathlib.Tactic.Linter.Header /-! ## Style linters -This file contains (currently one, eventually more) linters about stylistic aspects: -these are only about coding style, but do not affect correctness nor global coherence of mathlib. +This file contain linters about stylistic aspects: these are only about coding style, +but do not affect correctness nor global coherence of mathlib. +Historically, some of these were ported from the `lint-style.py` Python script. -Historically, these were ported from the `lint-style.py` Python script. +This file defines the following linters: +- the `set_option` linter checks for the presence of `set_option` commands activating +options disallowed in mathlib: these are meant to be temporary, and not for polished code +- the `missingEnd` linter checks for sections or namespaces which are not closed by the end +of the file: enforcing this invariant makes minimising files or moving code between files easier +- the `cdotLinter` linter checks for focusing dots `·` which are typed using a `.` instead: +this is allowed Lean syntax, but it is nicer to be uniform +- the `dollarSyntax` linter checks for use of the dollar sign `$` instead of the `<|` pipe operator: +similarly, both symbols have the same meaning, but mathlib prefers `<|` for the symmetry with +the `|>` symbol +- the `lambdaSyntax` linter checks for uses of the `λ` symbol for ananomous functions, +instead of the `fun` keyword: mathlib prefers the latter for reasons of readability +(This linter is still under review in PR #15896.) +- the `longLine` linter checks for lines which have more than 100 characters +- the `longFile` linter checks for files which have more than 1500 lines: +this linter is still under development in PR #15610. + +All of these linters are enabled in mathlib by default, but disabled globally +since they enforce conventions which are inherently subjective. -/ open Lean Elab Command @@ -71,4 +90,334 @@ initialize addLinter setOptionLinter end Style.setOption +/-! +# The "missing end" linter + +The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. +It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). +-/ + +open Lean Elab Command + +/-- The "missing end" linter emits a warning on non-closed `section`s and `namespace`s. +It allows the "outermost" `noncomputable section` to be left open (whether or not it is named). +-/ +register_option linter.style.missingEnd : Bool := { + defValue := false + descr := "enable the missing end linter" +} + +namespace Style.missingEnd + +@[inherit_doc Mathlib.Linter.linter.style.missingEnd] +def missingEndLinter : Linter where run := withSetOptionIn fun stx ↦ do + -- Only run this linter at the end of a module. + unless stx.isOfKind ``Lean.Parser.Command.eoi do return + if Linter.getLinterValue linter.style.missingEnd (← getOptions) && + !(← MonadState.get).messages.hasErrors then + let sc ← getScopes + -- The last scope is always the "base scope", corresponding to no active `section`s or + -- `namespace`s. We are interested in any *other* unclosed scopes. + if sc.length == 1 then return + let ends := sc.dropLast.map fun s ↦ (s.header, s.isNoncomputable) + -- If the outermost scope corresponds to a `noncomputable section`, we ignore it. + let ends := if ends.getLast!.2 then ends.dropLast else ends + -- If there are any further un-closed scopes, we emit a warning. + if !ends.isEmpty then + let ending := (ends.map Prod.fst).foldl (init := "") fun a b ↦ + a ++ s!"\n\nend{if b == "" then "" else " "}{b}" + Linter.logLint linter.style.missingEnd stx + m!"unclosed sections or namespaces; expected: '{ending}'" + +initialize addLinter missingEndLinter + +end Style.missingEnd + +/-! +# The `cdot` linter + +The `cdot` linter is a syntax-linter that flags uses of the "cdot" `·` that are achieved +by typing a character different from `·`. +For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. +It also flags "isolated cdots", i.e. when the `·` is on its own line. +-/ + +/-- +The `cdot` linter flags uses of the "cdot" `·` that are achieved by typing a character +different from `·`. +For instance, a "plain" dot `.` is allowed syntax, but is flagged by the linter. +It also flags "isolated cdots", i.e. when the `·` is on its own line. -/ +register_option linter.style.cdot : Bool := { + defValue := false + descr := "enable the `cdot` linter" +} + +/-- `isCDot? stx` checks whether `stx` is a `Syntax` node corresponding to a `cdot` typed with +the character `·`. -/ +def isCDot? : Syntax → Bool + | .node _ ``cdotTk #[.node _ `patternIgnore #[.node _ _ #[.atom _ v]]] => v == "·" + | .node _ ``Lean.Parser.Term.cdot #[.atom _ v] => v == "·" + | _ => false + +/-- +`findCDot stx` extracts from `stx` the syntax nodes of `kind` `Lean.Parser.Term.cdot` or `cdotTk`. -/ +partial +def findCDot : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findCDot).flatten + match kind with + | ``Lean.Parser.Term.cdot | ``cdotTk => dargs.push stx + | _ => dargs + |_ => #[] + +/-- `unwanted_cdot stx` returns an array of syntax atoms within `stx` +corresponding to `cdot`s that are not written with the character `·`. +This is precisely what the `cdot` linter flags. +-/ +def unwanted_cdot (stx : Syntax) : Array Syntax := + (findCDot stx).filter (!isCDot? ·) + +namespace Style + +@[inherit_doc linter.style.cdot] +def cdotLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.cdot (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in unwanted_cdot stx do + Linter.logLint linter.style.cdot s + m!"Please, use '·' (typed as `\\.`) instead of '{s}' as 'cdot'." + -- We also check for isolated cdot's, i.e. when the cdot is on its own line. + for cdot in Mathlib.Linter.findCDot stx do + match cdot.find? (·.isOfKind `token.«· ») with + | some (.node _ _ #[.atom (.original _ _ afterCDot _) _]) => + if (afterCDot.takeWhile (·.isWhitespace)).contains '\n' then + logWarningAt cdot <| .tagged linter.style.cdot.name + m!"This central dot `·` is isolated; please merge it with the next line." + | _ => return + +initialize addLinter cdotLinter + +end Style + +/-! +# The `dollarSyntax` linter + +The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. +These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. +-/ + +/-- The `dollarSyntax` linter flags uses of `<|` that are achieved by typing `$`. +These are disallowed by the mathlib style guide, as using `<|` pairs better with `|>`. -/ +register_option linter.style.dollarSyntax : Bool := { + defValue := false + descr := "enable the `dollarSyntax` linter" +} + +namespace Style.dollarSyntax + +/-- `findDollarSyntax stx` extracts from `stx` the syntax nodes of `kind` `$`. -/ +partial +def findDollarSyntax : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findDollarSyntax).flatten + match kind with + | ``«term_$__» => dargs.push stx + | _ => dargs + |_ => #[] + +@[inherit_doc linter.style.dollarSyntax] +def dollarSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.dollarSyntax (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in findDollarSyntax stx do + Linter.logLint linter.style.dollarSyntax s + m!"Please use '<|' instead of '$' for the pipe operator." + +initialize addLinter dollarSyntaxLinter + +end Style.dollarSyntax + +/-! +# The `lambdaSyntax` linter + +The `lambdaSyntax` linter is a syntax linter that flags uses of the symbol `λ` to define anonymous +functions, as opposed to the `fun` keyword. These are syntactically equivalent; mathlib style +prefers the latter as it is considered more readable. +-/ + +/-- +The `lambdaSyntax` linter flags uses of the symbol `λ` to define anonymous functions. +This is syntactically equivalent to the `fun` keyword; mathlib style prefers using the latter. +-/ +register_option linter.style.lambdaSyntax : Bool := { + defValue := false + descr := "enable the `lambdaSyntax` linter" +} + +namespace Style.lambdaSyntax + +/-- +`findLambdaSyntax stx` extracts from `stx` all syntax nodes of `kind` `Term.fun`. -/ +partial +def findLambdaSyntax : Syntax → Array Syntax + | stx@(.node _ kind args) => + let dargs := (args.map findLambdaSyntax).flatten + match kind with + | ``Parser.Term.fun => dargs.push stx + | _ => dargs + |_ => #[] + +@[inherit_doc linter.style.lambdaSyntax] +def lambdaSyntaxLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.lambdaSyntax (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + for s in findLambdaSyntax stx do + if let .atom _ "λ" := s[0] then + Linter.logLint linter.style.lambdaSyntax s[0] m!"\ + Please use 'fun' and not 'λ' to define anonymous functions.\n\ + The 'λ' syntax is deprecated in mathlib4." + +initialize addLinter lambdaSyntaxLinter + +end Style.lambdaSyntax + +/-! +# The "longFile" linter + +The "longFile" linter emits a warning on files which are longer than a certain number of lines +(1500 by default). +-/ + +/-- +The "longFile" linter emits a warning on files which are longer than a certain number of lines +(`linter.style.longFileDefValue` by default on mathlib, no limit for downstream projects). +If this option is set to `N` lines, the linter warns once a file has more than `N` lines. +A value of `0` silences the linter entirely. +-/ +register_option linter.style.longFile : Nat := { + defValue := 0 + descr := "enable the longFile linter" +} + +/-- The number of lines that the `longFile` linter considers the default. -/ +register_option linter.style.longFileDefValue : Nat := { + defValue := 1500 + descr := "a soft upper bound on the number of lines of each file" +} + +namespace Style.longFile + +@[inherit_doc Mathlib.Linter.linter.style.longFile] +def longFileLinter : Linter where run := withSetOptionIn fun stx ↦ do + let linterBound := linter.style.longFile.get (← getOptions) + if linterBound == 0 then + return + let defValue := linter.style.longFileDefValue.get (← getOptions) + let smallOption := match stx with + | `(set_option linter.style.longFile $x) => TSyntax.getNat ⟨x.raw⟩ ≤ defValue + | _ => false + if smallOption then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"The default value of the `longFile` linter is {defValue}.\n\ + The current value of {linterBound} does not exceed the allowed bound.\n\ + Please, remove the `set_option linter.style.longFile {linterBound}`." + else + -- Thanks to the above check, the linter option is either not set (and hence equal + -- to the default) or set to some value *larger* than the default. + -- `Parser.isTerminalCommand` allows `stx` to be `#exit`: this is useful for tests. + unless Parser.isTerminalCommand stx do return + -- We exclude `Mathlib.lean` from the linter: it exceeds linter's default number of allowed + -- lines, and it is an auto-generated import-only file. + -- TODO: if there are more such files, revise the implementation. + if (← getMainModule) == `Mathlib then return + if let some init := stx.getTailPos? then + -- the last line: we subtract 1, since the last line is expected to be empty + let lastLine := ((← getFileMap).toPosition init).line + -- In this case, the file has an allowed length, and the linter option is unnecessarily set. + if lastLine ≤ defValue && defValue < linterBound then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"The default value of the `longFile` linter is {defValue}.\n\ + This file is {lastLine} lines long which does not exceed the allowed bound.\n\ + Please, remove the `set_option linter.style.longFile {linterBound}`." + else + -- `candidate` is divisible by `100` and satisfies `lastLine + 100 < candidate ≤ lastLine + 200` + -- note that either `lastLine ≤ defValue` and `defValue = linterBound` hold or + -- `candidate` is necessarily bigger than `lastLine` and hence bigger than `defValue` + let candidate := (lastLine / 100) * 100 + 200 + let candidate := max candidate defValue + -- In this case, the file is longer than the default and also than what the option says. + if defValue ≤ linterBound && linterBound < lastLine then + logWarningAt stx <| .tagged linter.style.longFile.name + m!"This file is {lastLine} lines long, but the limit is {linterBound}.\n\n\ + You can extend the allowed length of the file using \ + `set_option linter.style.longFile {candidate}`.\n\ + You can completely disable this linter by setting the length limit to `0`." + else + -- Finally, the file exceeds the default value, but not the option: we only allow the value + -- of the option to be `candidate` or `candidate + 100`. + -- In particular, this flags any option that is set to an unnecessarily high value. + if linterBound == candidate || linterBound + 100 == candidate then return + else + logWarningAt stx <| .tagged linter.style.longFile.name + m!"This file is {lastLine} lines long. \ + The current limit is {linterBound}, but it is expected to be {candidate}:\n\ + `set_option linter.style.longFile {candidate}`." + +initialize addLinter longFileLinter + +end Style.longFile + +/-! # The "longLine linter" -/ + +/-- The "longLine" linter emits a warning on lines longer than 100 characters. +We allow lines containing URLs to be longer, though. -/ +register_option linter.style.longLine : Bool := { + defValue := false + descr := "enable the longLine linter" +} + +namespace Style.longLine + +@[inherit_doc Mathlib.Linter.linter.style.longLine] +def longLineLinter : Linter where run := withSetOptionIn fun stx ↦ do + unless Linter.getLinterValue linter.style.longLine (← getOptions) do + return + if (← MonadState.get).messages.hasErrors then + return + -- The linter ignores the `#guard_msgs` command, in particular its doc-string. + -- The linter still lints the message guarded by `#guard_msgs`. + if stx.isOfKind ``Lean.guardMsgsCmd then + return + -- if the linter reached the end of the file, then we scan the `import` syntax instead + let stx := ← do + if stx.isOfKind ``Lean.Parser.Command.eoi then + let fileMap ← getFileMap + -- `impMods` is the syntax for the modules imported in the current file + let (impMods, _) ← Parser.parseHeader + { input := fileMap.source, fileName := ← getFileName, fileMap := fileMap } + return impMods + else return stx + let sstr := stx.getSubstring? + let fm ← getFileMap + let longLines := ((sstr.getD default).splitOn "\n").filter fun line ↦ + (100 < (fm.toPosition line.stopPos).column) + for line in longLines do + if (line.splitOn "http").length ≤ 1 then + let stringMsg := if line.contains '"' then + "\nYou can use \"string gaps\" to format long strings: within a string quotation, \ + using a '\\' at the end of a line allows you to continue the string on the following \ + line, removing all intervening whitespace." + else "" + Linter.logLint linter.style.longLine (.ofRange ⟨line.startPos, line.stopPos⟩) + m!"This line exceeds the 100 character limit, please shorten it!{stringMsg}" +initialize addLinter longLineLinter + +end Style.longLine + end Mathlib.Linter diff --git a/MathlibTest/Lint.lean b/MathlibTest/Lint.lean index 1ab5b6b71ed53..3e03fc2dae24d 100644 --- a/MathlibTest/Lint.lean +++ b/MathlibTest/Lint.lean @@ -1,7 +1,5 @@ import Mathlib.Tactic.Linter.Lint import Mathlib.Tactic.ToAdditive -import Mathlib.Order.SetNotation - /-- warning: The namespace 'add' is duplicated in the declaration 'add.add' note: this linter can be disabled with `set_option linter.dupNamespace false` @@ -56,222 +54,3 @@ note: this linter can be disabled with `set_option linter.dupNamespace false` export Nat (add add_comm add) end add - -section cdotLinter -set_option linter.style.cdot true - -set_option linter.globalAttributeIn false in -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -attribute [instance] Int.add in -instance : Inhabited Nat where - default := by - . have := 0 - · have : Nat → Nat → Nat := (· + .) - . exact 0 - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -example : Add Nat where add := (. + ·) - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` --/ -#guard_msgs in -example : Add Nat where add := (. + ·) - -/-- -warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. -note: this linter can be disabled with `set_option linter.style.cdot false` ---- -warning: This central dot `·` is isolated; please merge it with the next line. ---- -warning: This central dot `·` is isolated; please merge it with the next line. --/ -#guard_msgs in -example : Nat := by - have : Nat := by - · - -- some empty have - have := 0 - · - - -- another - have := 1 - . exact 2 - exact 0 - -#guard_msgs in -example : True := by - have : Nat := by - -- This is how code should look: no error. - · -- comment - exact 37 - trivial - -end cdotLinter -set_option linter.style.dollarSyntax true - -set_option linter.globalAttributeIn false in -/-- -warning: Please use '<|' instead of '$' for the pipe operator. -note: this linter can be disabled with `set_option linter.style.dollarSyntax false` ---- -warning: Please use '<|' instead of '$' for the pipe operator. -note: this linter can be disabled with `set_option linter.style.dollarSyntax false` --/ -#guard_msgs in -attribute [instance] Int.add in -instance (f g : Nat → Nat) : Inhabited Nat where - default := by - · have := 0 - · have : Nat := f $ g $ 0 - · exact 0 - -section lambdaSyntaxLinter - -set_option linter.style.lambdaSyntax true - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -def foo : Bool := by - let _f : ℕ → ℕ := λ _ ↦ 0 - exact true - -example : ℕ → ℕ := fun n ↦ n - 1 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := by exact λ n ↦ 3 * n + 1 - -/-- -warning: declaration uses 'sorry' ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ → ℕ → ℕ := by - have (n : ℕ) : True := trivial - have : (Set.univ : Set ℕ) = ⋃ (i : ℕ), (Set.iUnion λ j ↦ ({0, j} : Set ℕ)) := sorry - have : ∃ m : ℕ, ⋃ i : ℕ, (Set.univ : Set ℕ) = ∅ := sorry - exact λ _a ↦ fun _b ↦ λ _c ↦ 0 - -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` ---- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : True := by - have : 0 = 0 ∧ 0 = 0 ∧ 1 + 3 = 4 := by - refine ⟨by trivial, by - let _f := λ n : ℕ ↦ 0; - have : ℕ := by - · -- comment - · have := λ k : ℕ ↦ -5 - · exact 0 - refine ⟨by trivial, have := λ k : ℕ ↦ -5; by simp⟩ - ⟩ - trivial - --- Code such as the following would require walking the infotree instead: --- the inner set_option is ignore (in either direction). --- As this seems unlikely to occur by accident and its use is dubious, we don't worry about this. -/-- -warning: Please use 'fun' and not 'λ' to define anonymous functions. -The 'λ' syntax is deprecated in mathlib4. -note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` --/ -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -set_option linter.style.lambdaSyntax false -#guard_msgs in -example : ℕ → ℕ := λ _ ↦ 0 - -end lambdaSyntaxLinter - -set_option linter.style.longLine true -/-- -warning: This line exceeds the 100 character limit, please shorten it! -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in -/-! -/ - -#guard_msgs in --- Lines with more than 100 characters containing URLs are allowed. -/-! http -/ - -set_option linter.style.longLine true --- The *argument* of `#guard_msgs` is *not* exempt from the linter. -/-- -warning: This line exceeds the 100 character limit, please shorten it! -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in #guard true - --- However, the *doc-string* of #guard_msgs is exempt from the linter: --- these are automatically generated, hence linting them is not helpful. -/-- -info: [0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26] --/ -#guard_msgs in -#eval List.range 27 - -/-- -info: " \" " : String ---- -warning: This line exceeds the 100 character limit, please shorten it! -You can use "string gaps" to format long strings: within a string quotation, using a '\' at the end of a line allows you to continue the string on the following line, removing all intervening whitespace. -note: this linter can be disabled with `set_option linter.style.longLine false` --/ -#guard_msgs in -#check " \" " diff --git a/MathlibTest/LintStyle.lean b/MathlibTest/LintStyle.lean index 82826979b6f8c..719fd833552a4 100644 --- a/MathlibTest/LintStyle.lean +++ b/MathlibTest/LintStyle.lean @@ -1,5 +1,5 @@ import Mathlib.Tactic.Linter.Style -import Mathlib.Tactic.Common +import Mathlib.Order.SetNotation /-! Tests for all the style linters. -/ @@ -120,3 +120,219 @@ lemma foo' : True := trivial -- TODO: add terms for the term form end setOption + +section cdotLinter +set_option linter.style.cdot true + +set_option linter.globalAttributeIn false in +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +attribute [instance] Int.add in +instance : Inhabited Nat where + default := by + . have := 0 + · have : Nat → Nat → Nat := (· + .) + . exact 0 + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +example : Add Nat where add := (. + ·) + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +-/ +#guard_msgs in +example : Add Nat where add := (. + ·) + +/-- +warning: Please, use '·' (typed as `\.`) instead of '.' as 'cdot'. +note: this linter can be disabled with `set_option linter.style.cdot false` +--- +warning: This central dot `·` is isolated; please merge it with the next line. +--- +warning: This central dot `·` is isolated; please merge it with the next line. +-/ +#guard_msgs in +example : Nat := by + have : Nat := by + · + -- some empty have + have := 0 + · + + -- another + have := 1 + . exact 2 + exact 0 + +#guard_msgs in +example : True := by + have : Nat := by + -- This is how code should look: no error. + · -- comment + exact 37 + trivial + +end cdotLinter +set_option linter.style.dollarSyntax true + +set_option linter.globalAttributeIn false in +/-- +warning: Please use '<|' instead of '$' for the pipe operator. +note: this linter can be disabled with `set_option linter.style.dollarSyntax false` +--- +warning: Please use '<|' instead of '$' for the pipe operator. +note: this linter can be disabled with `set_option linter.style.dollarSyntax false` +-/ +#guard_msgs in +attribute [instance] Int.add in +instance (f g : Nat → Nat) : Inhabited Nat where + default := by + · have := 0 + · have : Nat := f $ g $ 0 + · exact 0 + +section lambdaSyntaxLinter + +set_option linter.style.lambdaSyntax true + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +def foo : Bool := by + let _f : ℕ → ℕ := λ _ ↦ 0 + exact true + +example : ℕ → ℕ := fun n ↦ n - 1 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := by exact λ n ↦ 3 * n + 1 + +/-- +warning: declaration uses 'sorry' +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ → ℕ → ℕ := by + have (n : ℕ) : True := trivial + have : (Set.univ : Set ℕ) = ⋃ (i : ℕ), (Set.iUnion λ j ↦ ({0, j} : Set ℕ)) := sorry + have : ∃ m : ℕ, ⋃ i : ℕ, (Set.univ : Set ℕ) = ∅ := sorry + exact λ _a ↦ fun _b ↦ λ _c ↦ 0 + +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +--- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : True := by + have : 0 = 0 ∧ 0 = 0 ∧ 1 + 3 = 4 := by + refine ⟨by trivial, by + let _f := λ n : ℕ ↦ 0; + have : ℕ := by + · -- comment + · have := λ k : ℕ ↦ -5 + · exact 0 + refine ⟨by trivial, have := λ k : ℕ ↦ -5; by simp⟩ + ⟩ + trivial + +-- Code such as the following would require walking the infotree instead: +-- the inner set_option is ignore (in either direction). +-- As this seems unlikely to occur by accident and its use is dubious, we don't worry about this. +/-- +warning: Please use 'fun' and not 'λ' to define anonymous functions. +The 'λ' syntax is deprecated in mathlib4. +note: this linter can be disabled with `set_option linter.style.lambdaSyntax false` +-/ +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +set_option linter.style.lambdaSyntax false +#guard_msgs in +example : ℕ → ℕ := λ _ ↦ 0 + +end lambdaSyntaxLinter + +set_option linter.style.longLine true +/-- +warning: This line exceeds the 100 character limit, please shorten it! +note: this linter can be disabled with `set_option linter.style.longLine false` +-/ +#guard_msgs in +/-! -/ + +#guard_msgs in +-- Lines with more than 100 characters containing URLs are allowed. +/-! http -/ + +set_option linter.style.longLine true +-- The *argument* of `#guard_msgs` is *not* exempt from the linter. +/-- +warning: This line exceeds the 100 character limit, please shorten it! +note: this linter can be disabled with `set_option linter.style.longLine false` +-/ +#guard_msgs in #guard true + +-- However, the *doc-string* of #guard_msgs is exempt from the linter: +-- these are automatically generated, hence linting them is not helpful. +/-- +info: [0, 1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26] +-/ +#guard_msgs in +#eval List.range 27 + +-- TODO: this used to error about the 100 character limit (mentioning string gaps), +-- restore this test! +/-- info: " \" " : String -/ +#guard_msgs in +#check " \" \ + " diff --git a/MathlibTest/LongFile.lean b/MathlibTest/LongFile.lean index 635f2ab485c30..35b772bc6c409 100644 --- a/MathlibTest/LongFile.lean +++ b/MathlibTest/LongFile.lean @@ -1,4 +1,4 @@ -import Mathlib.Tactic.Linter.Lint +import Mathlib.Tactic.Linter.Style /- # Testing the `longFile` linter From 8489bbe7ee4cb7bea53236fda9a68368f16d5364 Mon Sep 17 00:00:00 2001 From: mathlib4-update-dependencies-bot <150093616+mathlib-bors@users.noreply.github.com> Date: Wed, 27 Nov 2024 12:53:27 +0000 Subject: [PATCH 231/250] chore: update Mathlib dependencies 2024-11-27 (#19536) 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 a3160f181a6cb..57203830de6b6 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "f6d16c2a073e0385a362ad3e5d9e7e8260e298d6", + "rev": "c933dd9b00271d869e22b802a015092d1e8e454a", "name": "batteries", "manifestFile": "lake-manifest.json", "inputRev": "main", From 37c90e46ca44c28b3c963d680ad3f4107279cbc1 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Wed, 27 Nov 2024 13:34:02 +0000 Subject: [PATCH 232/250] feat: uniform time lemma for the existence of global integral curves (#9013) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Lemma 9.15 in Lee's Introduction to Smooth Manifolds: > Let `v` be a smooth vector field on a smooth manifold `M`. If there exists `ε > 0` such that for each point `x : M`, there exists an integral curve of `v` through `x` defined on an open interval `Ioo (-ε) ε`, then every point on `M` has a global integral curve of `v` passing through it. We only require `v` to be $C^1$. To achieve this, we define the extension of an integral curve `γ` by another integral curve `γ'`, if they agree at a point inside their overlapping open interval domains. This utilises the uniqueness theorem of integral curves. We need this lemma to show that vector fields on compact manifolds always have global integral curves. - [x] depends on: #8886 - [x] depends on: #18833 - [x] depends on: #19008 Co-authored-by: Michael Rothgang Co-authored-by: Johan Commelin --- Mathlib.lean | 1 + .../Manifold/IntegralCurve/ExistUnique.lean | 4 +- .../Manifold/IntegralCurve/UniformTime.lean | 210 ++++++++++++++++++ Mathlib/Order/Interval/Set/Basic.lean | 8 +- 4 files changed, 218 insertions(+), 5 deletions(-) create mode 100644 Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean diff --git a/Mathlib.lean b/Mathlib.lean index a843c7987f61d..f3191d04fb72c 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -3047,6 +3047,7 @@ import Mathlib.Geometry.Manifold.Instances.UnitsOfNormedAlgebra import Mathlib.Geometry.Manifold.IntegralCurve.Basic import Mathlib.Geometry.Manifold.IntegralCurve.ExistUnique import Mathlib.Geometry.Manifold.IntegralCurve.Transform +import Mathlib.Geometry.Manifold.IntegralCurve.UniformTime import Mathlib.Geometry.Manifold.InteriorBoundary import Mathlib.Geometry.Manifold.LocalDiffeomorph import Mathlib.Geometry.Manifold.LocalInvariantProperties diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean index 369e3803fc7bd..8960304376097 100644 --- a/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean +++ b/Mathlib/Geometry/Manifold/IntegralCurve/ExistUnique.lean @@ -250,7 +250,7 @@ theorem isIntegralCurve_Ioo_eq_of_contMDiff_boundaryless [BoundarylessManifold I period `a - b`. -/ lemma IsIntegralCurve.periodic_of_eq [BoundarylessManifold I M] (hγ : IsIntegralCurve γ v) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (heq : γ a = γ b) : Periodic γ (a - b) := by intro t apply congrFun <| @@ -260,7 +260,7 @@ lemma IsIntegralCurve.periodic_of_eq [BoundarylessManifold I M] /-- A global integral curve is injective xor periodic with positive period. -/ lemma IsIntegralCurve.periodic_xor_injective [BoundarylessManifold I M] (hγ : IsIntegralCurve γ v) - (hv : ContMDiff I I.tangent 1 (fun x => (⟨x, v x⟩ : TangentBundle I M))) : + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) : Xor' (∃ T > 0, Periodic γ T) (Injective γ) := by rw [xor_iff_iff_not] refine ⟨fun ⟨T, hT, hf⟩ ↦ hf.not_injective (ne_of_gt hT), ?_⟩ diff --git a/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean b/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean new file mode 100644 index 0000000000000..6d11987df7df6 --- /dev/null +++ b/Mathlib/Geometry/Manifold/IntegralCurve/UniformTime.lean @@ -0,0 +1,210 @@ +/- +Copyright (c) 2023 Winston Yin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Winston Yin +-/ +import Mathlib.Geometry.Manifold.IntegralCurve.ExistUnique + +/-! +# Uniform time lemma for the global existence of integral curves + +## Main results + +* `exists_isIntegralCurve_of_isIntegralCurveOn`: If there exists `ε > 0` such that the local +integral curve at each point `x : M` is defined at least on an open interval `Ioo (-ε) ε`, then +every point on `M` has a global integral curve passing through it. + +## Reference + +* [Lee, J. M. (2012). _Introduction to Smooth Manifolds_. Springer New York.][lee2012] + +## Tags + +integral curve, vector field, global existence +-/ + +open scoped Topology + +open Function Set + +variable + {E : Type*} [NormedAddCommGroup E] [NormedSpace ℝ E] + {H : Type*} [TopologicalSpace H] {I : ModelWithCorners ℝ E H} + {M : Type*} [TopologicalSpace M] [ChartedSpace H M] [SmoothManifoldWithCorners I M] [T2Space M] + {γ γ' : ℝ → M} {v : (x : M) → TangentSpace I x} {s s' : Set ℝ} {t₀ : ℝ} + +/-- This is the uniqueness theorem of integral curves applied to a real-indexed family of integral + curves with the same starting point. -/ +lemma eqOn_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) + {a a' : ℝ} (hpos : 0 < a') (hle : a' ≤ a) : + EqOn (γ a') (γ a) (Ioo (-a') a') := by + apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless _ hv + (hγ a' (by positivity)) ((hγ a (gt_of_ge_of_gt hle hpos)).mono _) + (by rw [hγx a, hγx a']) + · rw [mem_Ioo] + exact ⟨neg_lt_zero.mpr hpos, by positivity⟩ + · apply Ioo_subset_Ioo <;> linarith + +/-- For a family of integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that + each `γ a` is defined on `Ioo (-a) a`, the global curve `γ_ext := fun t ↦ γ (|t| + 1) t` agrees + with each `γ a` on `Ioo (-a) a`. This will help us show that `γ_ext` is a global integral + curve. -/ +lemma eqOn_abs_add_one_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) + {a : ℝ} : EqOn (fun t ↦ γ (|t| + 1) t) (γ a) (Ioo (-a) a) := by + intros t ht + by_cases hlt : |t| + 1 < a + · exact eqOn_of_isIntegralCurveOn_Ioo hv γ hγx hγ + (by positivity) hlt.le (abs_lt.mp <| lt_add_one _) + · exact eqOn_of_isIntegralCurveOn_Ioo hv γ hγx hγ + (neg_lt_self_iff.mp <| lt_trans ht.1 ht.2) (not_lt.mp hlt) ht |>.symm + +/-- For a family of integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that + each `γ a` is defined on `Ioo (-a) a`, the function `γ_ext := fun t ↦ γ (|t| + 1) t` is a global + integral curve. -/ +lemma isIntegralCurve_abs_add_one_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) {x : M} + (γ : ℝ → ℝ → M) (hγx : ∀ a, γ a 0 = x) (hγ : ∀ a > 0, IsIntegralCurveOn (γ a) v (Ioo (-a) a)) : + IsIntegralCurve (fun t ↦ γ (|t| + 1) t) v := by + intro t + apply HasMFDerivAt.congr_of_eventuallyEq (f := γ (|t| + 1)) + · apply hγ (|t| + 1) (by positivity) + rw [mem_Ioo, ← abs_lt] + exact lt_add_one _ + · rw [Filter.eventuallyEq_iff_exists_mem] + refine ⟨Ioo (-(|t| + 1)) (|t| + 1), ?_, + eqOn_abs_add_one_of_isIntegralCurveOn_Ioo hv γ hγx hγ⟩ + have : |t| < |t| + 1 := lt_add_of_pos_right |t| zero_lt_one + rw [abs_lt] at this + exact Ioo_mem_nhds this.1 this.2 + +/-- The existence of a global integral curve is equivalent to the existence of a family of local + integral curves `γ : ℝ → ℝ → M` with the same starting point `γ 0 = x` such that each `γ a` is + defined on `Ioo (-a) a`. -/ +lemma exists_isIntegralCurve_iff_exists_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) (x : M) : + (∃ γ, γ 0 = x ∧ IsIntegralCurve γ v) ↔ + ∀ a, ∃ γ, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-a) a) := by + refine ⟨fun ⟨γ, h1, h2⟩ _ ↦ ⟨γ, h1, h2.isIntegralCurveOn _⟩, fun h ↦ ?_⟩ + choose γ hγx hγ using h + exact ⟨fun t ↦ γ (|t| + 1) t, hγx (|0| + 1), + isIntegralCurve_abs_add_one_of_isIntegralCurveOn_Ioo hv γ hγx (fun a _ ↦ hγ a)⟩ + +/-- Let `γ` and `γ'` be integral curves defined on `Ioo a b` and `Ioo a' b'`, respectively. Then, + `piecewise (Ioo a b) γ γ'` is equal to `γ` and `γ'` in their respective domains. + `Set.piecewise_eqOn` shows the equality for `γ` by definition, while this lemma shows the equality + for `γ'` by the uniqueness of integral curves. -/ +lemma eqOn_piecewise_of_isIntegralCurveOn_Ioo [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {a b a' b' : ℝ} (hγ : IsIntegralCurveOn γ v (Ioo a b)) + (hγ' : IsIntegralCurveOn γ' v (Ioo a' b')) + (ht₀ : t₀ ∈ Ioo a b ∩ Ioo a' b') (h : γ t₀ = γ' t₀) : + EqOn (piecewise (Ioo a b) γ γ') γ' (Ioo a' b') := by + intros t ht + suffices H : EqOn γ γ' (Ioo (max a a') (min b b')) by + by_cases hmem : t ∈ Ioo a b + · rw [piecewise, if_pos hmem] + apply H + simp [ht.1, ht.2, hmem.1, hmem.2] + · rw [piecewise, if_neg hmem] + apply isIntegralCurveOn_Ioo_eqOn_of_contMDiff_boundaryless _ hv + (hγ.mono (Ioo_subset_Ioo (le_max_left ..) (min_le_left ..))) + (hγ'.mono (Ioo_subset_Ioo (le_max_right ..) (min_le_right ..))) h + exact ⟨max_lt ht₀.1.1 ht₀.2.1, lt_min ht₀.1.2 ht₀.2.2⟩ + +/-- The extension of an integral curve by another integral curve is an integral curve. + + If two integral curves are defined on overlapping open intervals, and they agree at a point in + their common domain, then they can be patched together to form a longer integral curve. + + This is stated for manifolds without boundary for simplicity. We actually only need to assume that + the images of `γ` and `γ'` lie in the interior of the manifold. TODO: Generalise to manifolds with + boundary. -/ +lemma isIntegralCurveOn_piecewise [BoundarylessManifold I M] + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {a b a' b' : ℝ} (hγ : IsIntegralCurveOn γ v (Ioo a b)) + (hγ' : IsIntegralCurveOn γ' v (Ioo a' b')) {t₀ : ℝ} + (ht₀ : t₀ ∈ Ioo a b ∩ Ioo a' b') (h : γ t₀ = γ' t₀) : + IsIntegralCurveOn (piecewise (Ioo a b) γ γ') v (Ioo a b ∪ Ioo a' b') := by + intros t ht + by_cases hmem : t ∈ Ioo a b + · rw [piecewise, if_pos hmem] + apply (hγ t hmem).congr_of_eventuallyEq + rw [Filter.eventuallyEq_iff_exists_mem] + refine ⟨Ioo a b, isOpen_Ioo.mem_nhds hmem, ?_⟩ + intros t' ht' + rw [piecewise, if_pos ht'] + · rw [mem_union, or_iff_not_imp_left] at ht + rw [piecewise, if_neg hmem] + apply (hγ' t <| ht hmem).congr_of_eventuallyEq + rw [Filter.eventuallyEq_iff_exists_mem] + exact ⟨Ioo a' b', isOpen_Ioo.mem_nhds <| ht hmem, + eqOn_piecewise_of_isIntegralCurveOn_Ioo hv hγ hγ' ht₀ h⟩ + +/-- If there exists `ε > 0` such that the local integral curve at each point `x : M` is defined at + least on an open interval `Ioo (-ε) ε`, then every point on `M` has a global integral + curve passing through it. + + See Lemma 9.15, [J.M. Lee (2012)][lee2012]. -/ +lemma exists_isIntegralCurve_of_isIntegralCurveOn [BoundarylessManifold I M] + {v : (x : M) → TangentSpace I x} + (hv : ContMDiff I I.tangent 1 (fun x ↦ (⟨x, v x⟩ : TangentBundle I M))) + {ε : ℝ} (hε : 0 < ε) (h : ∀ x : M, ∃ γ : ℝ → M, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-ε) ε)) + (x : M) : ∃ γ : ℝ → M, γ 0 = x ∧ IsIntegralCurve γ v := by + let s := { a | ∃ γ, γ 0 = x ∧ IsIntegralCurveOn γ v (Ioo (-a) a) } + suffices hbdd : ¬BddAbove s by + rw [not_bddAbove_iff] at hbdd + rw [exists_isIntegralCurve_iff_exists_isIntegralCurveOn_Ioo hv] + intro a + obtain ⟨y, ⟨γ, hγ1, hγ2⟩, hlt⟩ := hbdd a + exact ⟨γ, hγ1, hγ2.mono <| Ioo_subset_Ioo (neg_le_neg hlt.le) hlt.le⟩ + intro hbdd + set asup := sSup s with hasup + -- we will obtain two integral curves, one centred at some `t₀ > 0` with + -- `0 ≤ asup - ε < t₀ < asup`; let `t₀ = asup - ε / 2` + -- another centred at 0 with domain up to `a ∈ S` with `t₀ < a < asup` + obtain ⟨a, ha, hlt⟩ := Real.add_neg_lt_sSup (⟨ε, h x⟩ : Set.Nonempty s) (ε := - (ε / 2)) + (by rw [neg_lt, neg_zero]; exact half_pos hε) + rw [mem_setOf] at ha + rw [← hasup, ← sub_eq_add_neg] at hlt + + -- integral curve defined on `Ioo (-a) a` + obtain ⟨γ, h0, hγ⟩ := ha + -- integral curve starting at `-(asup - ε / 2)` with radius `ε` + obtain ⟨γ1_aux, h1_aux, hγ1⟩ := h (γ (-(asup - ε / 2))) + rw [isIntegralCurveOn_comp_add (dt := asup - ε / 2)] at hγ1 + set γ1 := γ1_aux ∘ (· + (asup - ε / 2)) with γ1_def + have heq1 : γ1 (-(asup - ε / 2)) = γ (-(asup - ε / 2)) := by simp [γ1_def, h1_aux] + -- integral curve starting at `asup - ε / 2` with radius `ε` + obtain ⟨γ2_aux, h2_aux, hγ2⟩ := h (γ (asup - ε / 2)) + rw [isIntegralCurveOn_comp_sub (dt := asup - ε / 2)] at hγ2 + set γ2 := γ2_aux ∘ (· - (asup - ε / 2)) with γ2_def + have heq2 : γ2 (asup - ε / 2) = γ (asup - ε / 2) := by simp [γ2_def, h2_aux] + + -- rewrite shifted Ioo as Ioo + rw [neg_sub] at hγ1 + rw [Real.Ioo_eq_ball, neg_add_cancel, zero_div, sub_neg_eq_add, add_self_div_two, + Metric.vadd_ball, vadd_eq_add, add_zero, Real.ball_eq_Ioo] at hγ1 hγ2 + + -- to help `linarith` + have hεle : ε ≤ asup := le_csSup hbdd (h x) + + -- extend `γ` on the left by `γ1` and on the right by `γ2` + set γ_ext : ℝ → M := piecewise (Ioo (-(asup + ε / 2)) a) + (piecewise (Ioo (-a) a) γ γ1) γ2 with γ_ext_def + have heq_ext : γ_ext 0 = x := by + rw [γ_ext_def, piecewise, if_pos ⟨by linarith, by linarith⟩, piecewise, + if_pos ⟨by linarith, by linarith⟩, h0] + -- `asup + ε / 2` is an element of `s` greater than `asup`, a contradiction + suffices hext : IsIntegralCurveOn γ_ext v (Ioo (-(asup + ε / 2)) (asup + ε / 2)) from + (not_lt.mpr <| le_csSup hbdd ⟨γ_ext, heq_ext, hext⟩) <| lt_add_of_pos_right asup (half_pos hε) + apply (isIntegralCurveOn_piecewise (t₀ := asup - ε / 2) hv _ hγ2 + ⟨⟨by linarith, hlt⟩, ⟨by linarith, by linarith⟩⟩ + (by rw [piecewise, if_pos ⟨by linarith, hlt⟩, ← heq2])).mono + (Ioo_subset_Ioo_union_Ioo le_rfl (by linarith) (by linarith)) + exact (isIntegralCurveOn_piecewise (t₀ := -(asup - ε / 2)) hv hγ hγ1 + ⟨⟨neg_lt_neg hlt, by linarith⟩, ⟨by linarith, by linarith⟩⟩ heq1.symm).mono + (union_comm _ _ ▸ Ioo_subset_Ioo_union_Ioo (by linarith) (by linarith) le_rfl) diff --git a/Mathlib/Order/Interval/Set/Basic.lean b/Mathlib/Order/Interval/Set/Basic.lean index 7b2b1871c84e9..fb3aefe5d8181 100644 --- a/Mathlib/Order/Interval/Set/Basic.lean +++ b/Mathlib/Order/Interval/Set/Basic.lean @@ -1125,7 +1125,6 @@ theorem Iic_union_Ico_eq_Iio (h : a < b) : Iic a ∪ Ico a b = Iio b := /-! #### Two finite intervals, `I?o` and `Ic?` -/ - theorem Ioo_subset_Ioo_union_Ico : Ioo a c ⊆ Ioo a b ∪ Ico b c := fun x hx => (lt_or_le x b).elim (fun hxb => Or.inl ⟨hx.1, hxb⟩) fun hxb => Or.inr ⟨hxb, hx.2⟩ @@ -1181,7 +1180,6 @@ theorem Ioo_union_Icc_eq_Ioc (h₁ : a < b) (h₂ : b ≤ c) : Ioo a b ∪ Icc b /-! #### Two finite intervals, `I?c` and `Io?` -/ - theorem Ioo_subset_Ioc_union_Ioo : Ioo a c ⊆ Ioc a b ∪ Ioo b c := fun x hx => (le_or_lt x b).elim (fun hxb => Or.inl ⟨hx.1, hxb⟩) fun hxb => Or.inr ⟨hxb, hx.2⟩ @@ -1237,7 +1235,6 @@ theorem Ioc_union_Ioc (h₁ : min a b ≤ max c d) (h₂ : min c d ≤ max a b) /-! #### Two finite intervals with a common point -/ - theorem Ioo_subset_Ioc_union_Ico : Ioo a c ⊆ Ioc a b ∪ Ico b c := Subset.trans Ioo_subset_Ioc_union_Ioo (union_subset_union_right _ Ioo_subset_Ico_self) @@ -1317,6 +1314,11 @@ theorem Ioo_union_Ioo (h₁ : min a b < max c d) (h₂ : min c d < max a b) : simp [*, min_eq_left_of_lt, min_eq_right_of_lt, max_eq_left_of_lt, max_eq_right_of_lt, le_of_lt h₂, le_of_lt h₁] +theorem Ioo_subset_Ioo_union_Ioo (h₁ : a ≤ a₁) (h₂ : c < b) (h₃ : b₁ ≤ d) : + Ioo a₁ b₁ ⊆ Ioo a b ∪ Ioo c d := fun x hx => + (lt_or_le x b).elim (fun hxb => Or.inl ⟨lt_of_le_of_lt h₁ hx.1, hxb⟩) + fun hxb => Or.inr ⟨lt_of_lt_of_le h₂ hxb, lt_of_lt_of_le hx.2 h₃⟩ + end LinearOrder section Lattice From af58035408a3fc62e9ad7e709a01d4f5f97cd714 Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 14:12:32 +0000 Subject: [PATCH 233/250] chore(discover-lean-pr-testing): add PR number and diff link to relevant branches (#19532) --- .github/workflows/discover-lean-pr-testing.yml | 9 +++++++-- 1 file changed, 7 insertions(+), 2 deletions(-) diff --git a/.github/workflows/discover-lean-pr-testing.yml b/.github/workflows/discover-lean-pr-testing.yml index 349945a0ee8fe..fa3dcad9ae5fe 100644 --- a/.github/workflows/discover-lean-pr-testing.yml +++ b/.github/workflows/discover-lean-pr-testing.yml @@ -100,8 +100,13 @@ jobs: # Check if the diff contains files other than the specified ones if echo "$DIFF_FILES" | grep -v -e 'lake-manifest.json' -e 'lakefile.lean' -e 'lean-toolchain'; then - # If it does, add the branch to RELEVANT_BRANCHES - RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$BRANCH" + # Extract the PR number and actual branch name + PR_NUMBER=$(echo "$BRANCH" | grep -oP '\d+$') + ACTUAL_BRANCH=${BRANCH#origin/} + GITHUB_DIFF="https://github.com/leanprover-community/mathlib4/compare/nightly-testing...$ACTUAL_BRANCH" + + # Append the branch details to RELEVANT_BRANCHES + RELEVANT_BRANCHES="$RELEVANT_BRANCHES"$'\n- '"$ACTUAL_BRANCH (lean4#$PR_NUMBER) ([diff with \`nightly-testing\`]($GITHUB_DIFF))" fi done From 8390aa74f64d70fa884b2591c9d3bf9e54ebb3fe Mon Sep 17 00:00:00 2001 From: Johan Commelin Date: Wed, 27 Nov 2024 14:12:34 +0000 Subject: [PATCH 234/250] chore(zulip_emoji_merge_delegate): exclude rss channel in python logic instead of zulip search (#19540) --- scripts/zulip_emoji_merge_delegate.py | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/scripts/zulip_emoji_merge_delegate.py b/scripts/zulip_emoji_merge_delegate.py index 91836a41e49ff..f31e66c7b506f 100644 --- a/scripts/zulip_emoji_merge_delegate.py +++ b/scripts/zulip_emoji_merge_delegate.py @@ -23,14 +23,14 @@ site=ZULIP_SITE ) -# Fetch the messages containing the PR number from the public channels (exluding #rss). +# Fetch the messages containing the PR number from the public channels. # There does not seem to be a way to search simultaneously public and private channels. public_response = client.get_messages({ "anchor": "newest", "num_before": 5000, "num_after": 0, "narrow": [ - {"operator": "channel", "operand": "rss", "negated": True}, + {"operator": "channels", "operand": "public"}, {"operator": "search", "operand": f'#{PR_NUMBER}'}, ], }) @@ -57,6 +57,8 @@ print(f"Searching for: '{pr_pattern}'") for message in messages: + if message['display_recipient'] == 'rss': + continue content = message['content'] # Check for emoji reactions reactions = message['reactions'] From b5fb043af905e8b3bb5837dafc9b0b55d12e7092 Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Wed, 27 Nov 2024 15:22:44 +0000 Subject: [PATCH 235/250] feat(Fin/Tuple): add `Fin.update_insertNth` (#19542) --- Mathlib/Data/Fin/Tuple/Basic.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/Mathlib/Data/Fin/Tuple/Basic.lean b/Mathlib/Data/Fin/Tuple/Basic.lean index dd1645bcb7380..430153e9bdb75 100644 --- a/Mathlib/Data/Fin/Tuple/Basic.lean +++ b/Mathlib/Data/Fin/Tuple/Basic.lean @@ -949,6 +949,12 @@ open Set lemma insertNth_self_removeNth (p : Fin (n + 1)) (f : ∀ j, α j) : insertNth p (f p) (removeNth p f) = f := by simp +@[simp] +theorem update_insertNth (p : Fin (n + 1)) (x y : α p) (f : ∀ i, α (p.succAbove i)) : + update (p.insertNth x f) p y = p.insertNth y f := by + ext i + cases i using p.succAboveCases <;> simp [succAbove_ne] + /-- Equivalence between tuples of length `n + 1` and pairs of an element and a tuple of length `n` given by separating out the `p`-th element of the tuple. From 6a0cdd15b083f05786bf1de0eb7df4c57d925d6e Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Jo=C3=ABl=20Riou?= Date: Wed, 27 Nov 2024 18:07:36 +0000 Subject: [PATCH 236/250] feat(Algebra/Module): a presentation of an algebra induces a presentation of the module of differentials (#18440) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit Co-authored-by: Joël Riou <37772949+joelriou@users.noreply.github.com> --- Mathlib.lean | 1 + .../Module/Presentation/Differentials.lean | 164 ++++++++++++++++++ 2 files changed, 165 insertions(+) create mode 100644 Mathlib/Algebra/Module/Presentation/Differentials.lean diff --git a/Mathlib.lean b/Mathlib.lean index f3191d04fb72c..589e1a973bc7e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -545,6 +545,7 @@ import Mathlib.Algebra.Module.Pi import Mathlib.Algebra.Module.PointwisePi import Mathlib.Algebra.Module.Presentation.Basic import Mathlib.Algebra.Module.Presentation.Cokernel +import Mathlib.Algebra.Module.Presentation.Differentials import Mathlib.Algebra.Module.Presentation.DirectSum import Mathlib.Algebra.Module.Presentation.Finite import Mathlib.Algebra.Module.Presentation.Free diff --git a/Mathlib/Algebra/Module/Presentation/Differentials.lean b/Mathlib/Algebra/Module/Presentation/Differentials.lean new file mode 100644 index 0000000000000..5ec008c5b1ed6 --- /dev/null +++ b/Mathlib/Algebra/Module/Presentation/Differentials.lean @@ -0,0 +1,164 @@ +/- +Copyright (c) 2024 Joël Riou. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joël Riou +-/ + +import Mathlib.Algebra.Module.Presentation.Basic +import Mathlib.RingTheory.Kaehler.Polynomial +import Mathlib.RingTheory.Kaehler.CotangentComplex +import Mathlib.RingTheory.Presentation + +/-! +# Presentation of the module of differentials + +Given a presentation of a `R`-algebra `S`, we obtain a presentation +of the `S`-module `Ω[S⁄R]`. + +Assume `pres : Algebra.Presentation R S` is a presentation of `S` as an `R`-algebra. +We then have a type `pres.vars` for the generators, a type `pres.rels` for the relations, +and for each `r : pres.rels` we have the relation `pres.relation r` in `pres.Ring` (which is +a ring of polynomials over `R` with variables indexed by `pres.vars`). +Then, `Ω[pres.Ring⁄R]` identifies to the free module on the type `pres.vars`, and +for each `r : pres.rels`, we may consider the image of the differential of `pres.relation r` +in this free module, and by using the (surjective) map `pres.Ring → S`, we obtain +an element `pres.differentialsRelations.relation r` in the free `S`-module on `pres.vars`. +The main result in this file is that `Ω[S⁄R]` identifies to the quotient of the +free `S`-module on `pres.vars` by the submodule generated by these +elements `pres.differentialsRelations.relation r`. We show this as a consequence +of `Algebra.Extension.exact_cotangentComplex_toKaehler` +from the file `Mathlib.RingTheory.Kaehler.CotangentComplex`. + +-/ + +universe w' t w u v + +namespace Algebra.Presentation + +open KaehlerDifferential + +variable {R : Type u} {S : Type v} [CommRing R] [CommRing S] [Algebra R S] + (pres : Algebra.Presentation.{t, w} R S) + +/-- The shape of the presentation by generators and relations of the `S`-module `Ω[S⁄R]` +that is obtained from a presentation of `S` as an `R`-algebra. -/ +@[simps G R] +noncomputable def differentialsRelations : Module.Relations S where + G := pres.vars + R := pres.rels + relation r := + Finsupp.mapRange (algebraMap pres.Ring S) (by simp) + ((mvPolynomialBasis R pres.vars).repr (D _ _ (pres.relation r))) + +namespace differentials + +/-! We obtain here a few compatibilities which allow to compare the two sequences +`(pres.rels →₀ S) → (pres.vars →₀ S) → Ω[S⁄R]` and +`pres.toExtension.Cotangent →ₗ[S] pres.toExtension.CotangentSpace → Ω[S⁄R]`. +Indeed, there is commutative diagram with a surjective map +`hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent` on the left and +bijections on the middle and on the right. Then, the exactness of the first +sequence shall follow from the exactness of the second which is +`Algebra.Extension.exact_cotangentComplex_toKaehler`. -/ + +/-- Same as `comm₂₃` below, but here we have not yet constructed `differentialsSolution`. -/ +lemma comm₂₃' : pres.toExtension.toKaehler.comp pres.cotangentSpaceBasis.repr.symm.toLinearMap = + Finsupp.linearCombination S (fun g ↦ D _ _ (pres.val g)) := by + ext g + dsimp + rw [Basis.repr_symm_apply, Finsupp.linearCombination_single, + Finsupp.linearCombination_single, one_smul, one_smul, + Generators.cotangentSpaceBasis_apply, mapBaseChange_tmul, one_smul] + simp + +/-- The canonical map `(pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent`. -/ +noncomputable def hom₁ : (pres.rels →₀ S) →ₗ[S] pres.toExtension.Cotangent := + Finsupp.linearCombination S (fun r ↦ Extension.Cotangent.mk ⟨pres.relation r, by simp⟩) + +lemma hom₁_single (r : pres.rels) : + hom₁ pres (Finsupp.single r 1) = Extension.Cotangent.mk ⟨pres.relation r, by simp⟩ := by + simp [hom₁] + +lemma surjective_hom₁ : Function.Surjective (hom₁ pres) := by + let φ : (pres.rels →₀ S) →ₗ[pres.Ring] pres.toExtension.Cotangent := + { toFun := hom₁ pres + map_add' := by simp + map_smul' := by simp } + change Function.Surjective φ + have h₁ := Algebra.Extension.Cotangent.mk_surjective (P := pres.toExtension) + have h₂ : Submodule.span pres.Ring + (Set.range (fun r ↦ (⟨pres.relation r, by simp⟩ : pres.ker))) = ⊤ := by + refine Submodule.map_injective_of_injective (f := Submodule.subtype pres.ker) + Subtype.coe_injective ?_ + rw [Submodule.map_top, Submodule.range_subtype, Submodule.map_span, + Submodule.coe_subtype, Ideal.submodule_span_eq] + simp only [← pres.span_range_relation_eq_ker] + congr + aesop + rw [← LinearMap.range_eq_top] at h₁ ⊢ + rw [← top_le_iff, ← h₁, LinearMap.range_eq_map, ← h₂] + dsimp + rw [Submodule.map_span_le] + rintro _ ⟨r, rfl⟩ + simp only [LinearMap.mem_range] + refine ⟨Finsupp.single r 1, ?_⟩ + simp only [LinearMap.coe_mk, AddHom.coe_mk, hom₁_single, φ] + rfl + +lemma comm₁₂_single (r : pres.rels) : + pres.toExtension.cotangentComplex (hom₁ pres (Finsupp.single r 1)) = + pres.cotangentSpaceBasis.repr.symm ((differentialsRelations pres).relation r) := by + simp only [hom₁, Finsupp.linearCombination_single, one_smul, differentialsRelations, + Basis.repr_symm_apply, Extension.cotangentComplex_mk] + exact pres.cotangentSpaceBasis.repr.injective (by ext; simp) + +lemma comm₁₂ : pres.toExtension.cotangentComplex.comp (hom₁ pres) = + pres.cotangentSpaceBasis.repr.symm.comp (differentialsRelations pres).map := by + ext r + have := (differentialsRelations pres).map_single + dsimp at this ⊢ + rw [comm₁₂_single, this] + +end differentials + +open differentials in +/-- The `S`-module `Ω[S⁄R]` contains an obvious solution to the system of linear +equations `pres.differentialsRelations.Solution` when `pres` is a presentation +of `S` as an `R`-algebra. -/ +noncomputable def differentialsSolution : + pres.differentialsRelations.Solution (Ω[S⁄R]) where + var g := D _ _ (pres.val g) + linearCombination_var_relation r := by + simp only [differentialsRelations_G, LinearMap.coe_comp, LinearEquiv.coe_coe, + Function.comp_apply, ← comm₂₃', ← comm₁₂_single] + apply DFunLike.congr_fun (Function.Exact.linearMap_comp_eq_zero + (pres.toExtension.exact_cotangentComplex_toKaehler)) + +lemma differentials.comm₂₃ : + pres.toExtension.toKaehler.comp pres.cotangentSpaceBasis.repr.symm.toLinearMap = + pres.differentialsSolution.π := + comm₂₃' pres + +open differentials in +lemma differentialsSolution_isPresentation : + pres.differentialsSolution.IsPresentation := by + rw [Module.Relations.Solution.isPresentation_iff] + constructor + · rw [← Module.Relations.Solution.surjective_π_iff_span_eq_top, ← comm₂₃] + exact Extension.toKaehler_surjective.comp pres.cotangentSpaceBasis.repr.symm.surjective + · rw [← Module.Relations.range_map] + exact Function.Exact.linearMap_ker_eq + ((LinearMap.exact_iff_of_surjective_of_bijective_of_injective + _ _ _ _ (hom₁ pres) + pres.cotangentSpaceBasis.repr.symm.toLinearMap .id + (comm₁₂ pres) (by simpa using comm₂₃ pres) (surjective_hom₁ pres) + (LinearEquiv.bijective _) (Equiv.refl _).injective).2 + pres.toExtension.exact_cotangentComplex_toKaehler) + +/-- The presentation of the `S`-module `Ω[S⁄R]` deduced from a presentation +of `S` as a `R`-algebra. -/ +noncomputable def differentials : Module.Presentation S (Ω[S⁄R]) where + toSolution := differentialsSolution pres + toIsPresentation := pres.differentialsSolution_isPresentation + +end Algebra.Presentation From db422a292b959d8dec7f19c72494ba4915869d78 Mon Sep 17 00:00:00 2001 From: Emily Riehl Date: Wed, 27 Nov 2024 18:16:39 +0000 Subject: [PATCH 237/250] feat(CategoryTheory): coskeletal simplicial objects (#19492) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit A simplicial object `X` is *n-coskeletal* if the identity natural transformation `X` is a right extension of its restriction along `(Truncated.inclusion (n := n)).op` recorded by `rightExtensionInclusion X n`. Co-Authored-By: [Mario Carneiro](https://github.com/digama0) and [Joël Riou](https://github.com/joelriou) --- Mathlib.lean | 3 +- Mathlib/AlgebraicTopology/CechNerve.lean | 2 +- Mathlib/AlgebraicTopology/MooreComplex.lean | 2 +- .../AlgebraicTopology/SimplexCategory.lean | 8 +- .../Basic.lean} | 38 +++++--- .../SimplicialObject/Coskeletal.lean | 91 +++++++++++++++++++ .../SimplicialSet/Basic.lean | 2 +- .../SplitSimplicialObject.lean | 2 +- .../Idempotents/SimplicialObject.lean | 2 +- 9 files changed, 126 insertions(+), 24 deletions(-) rename Mathlib/AlgebraicTopology/{SimplicialObject.lean => SimplicialObject/Basic.lean} (96%) create mode 100644 Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean diff --git a/Mathlib.lean b/Mathlib.lean index 589e1a973bc7e..4b3dcced1bc36 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -1042,7 +1042,8 @@ import Mathlib.AlgebraicTopology.Quasicategory.StrictSegal import Mathlib.AlgebraicTopology.SimplexCategory import Mathlib.AlgebraicTopology.SimplicialCategory.Basic import Mathlib.AlgebraicTopology.SimplicialCategory.SimplicialObject -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic +import Mathlib.AlgebraicTopology.SimplicialObject.Coskeletal import Mathlib.AlgebraicTopology.SimplicialSet.Basic import Mathlib.AlgebraicTopology.SimplicialSet.KanComplex import Mathlib.AlgebraicTopology.SimplicialSet.Monoidal diff --git a/Mathlib/AlgebraicTopology/CechNerve.lean b/Mathlib/AlgebraicTopology/CechNerve.lean index 0142017c6b6ab..667421531fc21 100644 --- a/Mathlib/AlgebraicTopology/CechNerve.lean +++ b/Mathlib/AlgebraicTopology/CechNerve.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Adam Topaz. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Adam Topaz -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Comma.Arrow import Mathlib.CategoryTheory.Limits.Shapes.WidePullbacks import Mathlib.CategoryTheory.Limits.Shapes.FiniteProducts diff --git a/Mathlib/AlgebraicTopology/MooreComplex.lean b/Mathlib/AlgebraicTopology/MooreComplex.lean index 0ce1710a542c8..0276dc0ce398a 100644 --- a/Mathlib/AlgebraicTopology/MooreComplex.lean +++ b/Mathlib/AlgebraicTopology/MooreComplex.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kim Morrison -/ import Mathlib.Algebra.Homology.HomologicalComplex -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Abelian.Basic /-! diff --git a/Mathlib/AlgebraicTopology/SimplexCategory.lean b/Mathlib/AlgebraicTopology/SimplexCategory.lean index 7c9ce0278a0c3..115b2bc0a1851 100644 --- a/Mathlib/AlgebraicTopology/SimplexCategory.lean +++ b/Mathlib/AlgebraicTopology/SimplexCategory.lean @@ -725,15 +725,15 @@ instance {n} : Inhabited (Truncated n) := /-- The fully faithful inclusion of the truncated simplex category into the usual simplex category. -/ -def inclusion {n : ℕ} : SimplexCategory.Truncated n ⥤ SimplexCategory := +def inclusion (n : ℕ) : SimplexCategory.Truncated n ⥤ SimplexCategory := fullSubcategoryInclusion _ -instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Full := FullSubcategory.full _ -instance (n : ℕ) : (inclusion : Truncated n ⥤ _).Faithful := FullSubcategory.faithful _ +instance (n : ℕ) : (inclusion n : Truncated n ⥤ _).Full := FullSubcategory.full _ +instance (n : ℕ) : (inclusion n : Truncated n ⥤ _).Faithful := FullSubcategory.faithful _ /-- A proof that the full subcategory inclusion is fully faithful.-/ noncomputable def inclusion.fullyFaithful (n : ℕ) : - (inclusion : Truncated n ⥤ _).op.FullyFaithful := Functor.FullyFaithful.ofFullyFaithful _ + (inclusion n : Truncated n ⥤ _).op.FullyFaithful := Functor.FullyFaithful.ofFullyFaithful _ @[ext] theorem Hom.ext {n} {a b : Truncated n} (f g : a ⟶ b) : diff --git a/Mathlib/AlgebraicTopology/SimplicialObject.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean similarity index 96% rename from Mathlib/AlgebraicTopology/SimplicialObject.lean rename to Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean index a7a5cb0b002e5..c7ee1ad189a72 100644 --- a/Mathlib/AlgebraicTopology/SimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Basic.lean @@ -250,7 +250,7 @@ section Truncation /-- The truncation functor from simplicial objects to truncated simplicial objects. -/ def truncation (n : ℕ) : SimplicialObject C ⥤ SimplicialObject.Truncated C n := - (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion.op + (whiskeringLeft _ _ _).obj (SimplexCategory.Truncated.inclusion n).op end Truncation @@ -259,24 +259,24 @@ noncomputable section /-- The n-skeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ protected abbrev Truncated.sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] : SimplicialObject.Truncated C n ⥤ SimplicialObject C := - lan (SimplexCategory.Truncated.inclusion.op) + lan (SimplexCategory.Truncated.inclusion n).op /-- The n-coskeleton as a functor `SimplicialObject.Truncated C n ⥤ SimplicialObject C`. -/ protected abbrev Truncated.cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] : SimplicialObject.Truncated C n ⥤ SimplicialObject C := - ran (SimplexCategory.Truncated.inclusion.op) + ran (SimplexCategory.Truncated.inclusion n).op /-- The n-skeleton as an endofunctor on `SimplicialObject C`. -/ abbrev sk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] : SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.sk n /-- The n-coskeleton as an endofunctor on `SimplicialObject C`. -/ abbrev cosk (n : ℕ) [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] : + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] : SimplicialObject C ⥤ SimplicialObject C := truncation n ⋙ Truncated.cosk n end @@ -288,9 +288,9 @@ respectively define left and right adjoints to `truncation n`.-/ variable (n : ℕ) variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasRightKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasLeftKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasLeftKanExtension F] /-- The adjunction between the n-skeleton and n-truncation.-/ noncomputable def skAdj : Truncated.sk (C := C) n ⊣ truncation n := @@ -300,20 +300,30 @@ noncomputable def skAdj : Truncated.sk (C := C) n ⊣ truncation n := noncomputable def coskAdj : truncation (C := C) n ⊣ Truncated.cosk n := ranAdjunction _ _ +instance : ((sk n).obj X).IsLeftKanExtension ((skAdj n).unit.app _) := by + dsimp [sk, skAdj] + rw [lanAdjunction_unit] + infer_instance + +instance : ((cosk n).obj X).IsRightKanExtension ((coskAdj n).counit.app _) := by + dsimp [cosk, coskAdj] + rw [ranAdjunction_counit] + infer_instance + namespace Truncated /- When the left and right Kan extensions exist and are pointwise Kan extensions, `skAdj n` and `coskAdj n` are respectively coreflective and reflective.-/ variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasPointwiseRightKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasPointwiseRightKanExtension F] variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), - SimplexCategory.Truncated.inclusion.op.HasPointwiseLeftKanExtension F] + (SimplexCategory.Truncated.inclusion n).op.HasPointwiseLeftKanExtension F] instance cosk_reflective : IsIso (coskAdj (C := C) n).counit := - reflective' SimplexCategory.Truncated.inclusion.op + reflective' (SimplexCategory.Truncated.inclusion n).op instance sk_coreflective : IsIso (skAdj (C := C) n).unit := - coreflective' SimplexCategory.Truncated.inclusion.op + coreflective' (SimplexCategory.Truncated.inclusion n).op /-- Since `Truncated.inclusion` is fully faithful, so is right Kan extension along it.-/ noncomputable def cosk.fullyFaithful : @@ -675,7 +685,7 @@ section Truncation /-- The truncation functor from cosimplicial objects to truncated cosimplicial objects. -/ def truncation (n : ℕ) : CosimplicialObject C ⥤ CosimplicialObject.Truncated C n := - (whiskeringLeft _ _ _).obj SimplexCategory.Truncated.inclusion + (whiskeringLeft _ _ _).obj (SimplexCategory.Truncated.inclusion n) end Truncation diff --git a/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean b/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean new file mode 100644 index 0000000000000..247afdd825990 --- /dev/null +++ b/Mathlib/AlgebraicTopology/SimplicialObject/Coskeletal.lean @@ -0,0 +1,91 @@ +/- +Copyright (c) 2024 Emily Riehl. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro, Emily Riehl, Joël Riou +-/ +import Mathlib.AlgebraicTopology.SimplicialObject.Basic +import Mathlib.CategoryTheory.Functor.KanExtension.Adjunction +import Mathlib.CategoryTheory.Functor.KanExtension.Basic + +/-! +# Coskeletal simplicial objects + +The identity natural transformation exhibits a simplicial object `X` as a right extension of its +restriction along `(Truncated.inclusion n).op` recorded by `rightExtensionInclusion X n`. + +The simplicial object `X` is *n-coskeletal* if `rightExtensionInclusion X n` is a right Kan +extension. + +When the ambient category admits right Kan extensions along `(Truncated.inclusion n).op`, +then when `X` is `n`-coskeletal, the unit of `coskAdj n` defines an isomorphism: +`isoCoskOfIsCoskeletal : X ≅ (cosk n).obj X`. + +TODO: Prove that `X` is `n`-coskeletal whenever a certain canonical cone is a limit cone. +-/ + +open Opposite + +open CategoryTheory + +open CategoryTheory.Limits CategoryTheory.Functor SimplexCategory + +universe v u v' u' + +namespace CategoryTheory + +namespace SimplicialObject +variable {C : Type u} [Category.{v} C] +variable (X : SimplicialObject C) (n : ℕ) + +namespace Truncated + +/-- The identity natural transformation exhibits a simplicial set as a right extension of its +restriction along `(Truncated.inclusion n).op`.-/ +@[simps!] +def rightExtensionInclusion : + RightExtension (Truncated.inclusion n).op + ((Truncated.inclusion n).op ⋙ X) := RightExtension.mk _ (𝟙 _) + +end Truncated + +open Truncated + +/-- A simplicial object `X` is `n`-coskeletal when it is the right Kan extension of its restriction +along `(Truncated.inclusion n).op` via the identity natural transformation. -/ +@[mk_iff] +class IsCoskeletal : Prop where + isRightKanExtension : IsRightKanExtension X (𝟙 ((Truncated.inclusion n).op ⋙ X)) + +attribute [instance] IsCoskeletal.isRightKanExtension + +section + +variable [∀ (F : (SimplexCategory.Truncated n)ᵒᵖ ⥤ C), + (SimplexCategory.Truncated.inclusion n).op.HasRightKanExtension F] + +/-- If `X` is `n`-coskeletal, then `Truncated.rightExtensionInclusion X n` is a terminal object in +the category `RightExtension (Truncated.inclusion n).op (Truncated.inclusion.op ⋙ X)`. -/ +noncomputable def IsCoskeletal.isUniversalOfIsRightKanExtension [X.IsCoskeletal n] : + (rightExtensionInclusion X n).IsUniversal := by + apply Functor.isUniversalOfIsRightKanExtension + +theorem isCoskeletal_iff_isIso : X.IsCoskeletal n ↔ IsIso ((coskAdj n).unit.app X) := by + rw [isCoskeletal_iff] + exact isRightKanExtension_iff_isIso ((coskAdj n).unit.app X) + ((coskAdj n).counit.app _) (𝟙 _) ((coskAdj n).left_triangle_components X) + +instance [X.IsCoskeletal n] : IsIso ((coskAdj n).unit.app X) := by + rw [← isCoskeletal_iff_isIso] + infer_instance + +/-- The canonical isomorphism `X ≅ (cosk n).obj X` defined when `X` is coskeletal and the +`n`-coskeleton functor exists.-/ +@[simps! hom] +noncomputable def isoCoskOfIsCoskeletal [X.IsCoskeletal n] : X ≅ (cosk n).obj X := + asIso ((coskAdj n).unit.app X) + +end + +end SimplicialObject + +end CategoryTheory diff --git a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean index 8c329680dec88..6e545ab1fa48e 100644 --- a/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean +++ b/Mathlib/AlgebraicTopology/SimplicialSet/Basic.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Kim Morrison. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Kim Morrison, Adam Topaz -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Limits.Shapes.Types import Mathlib.CategoryTheory.Yoneda import Mathlib.Data.Fin.VecNotation diff --git a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean index db4c3e1b3959c..7e6de96ffecde 100644 --- a/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean +++ b/Mathlib/AlgebraicTopology/SplitSimplicialObject.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Limits.Shapes.Products import Mathlib.Data.Fintype.Sigma diff --git a/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean b/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean index bcccaf7a24640..0a46ade189089 100644 --- a/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean +++ b/Mathlib/CategoryTheory/Idempotents/SimplicialObject.lean @@ -3,7 +3,7 @@ Copyright (c) 2022 Joël Riou. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joël Riou -/ -import Mathlib.AlgebraicTopology.SimplicialObject +import Mathlib.AlgebraicTopology.SimplicialObject.Basic import Mathlib.CategoryTheory.Idempotents.FunctorCategories /-! From 0aebe464b32804553b79f0268c4030e3d8219a1c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 19:53:47 +0000 Subject: [PATCH 238/250] refactor(Combinatorics/Additive): generalise Ruzsa's covering lemma to non-abelian groups (#19512) Also state it using a fixed real number and use semantic lemma names. This is much more useful in practice. From GrowthInGroups (LeanCamCombi) --- .../Combinatorics/Additive/RuzsaCovering.lean | 83 ++++++++++--------- 1 file changed, 43 insertions(+), 40 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 6e7183a4afdf7..825ab42bd37ab 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -4,68 +4,71 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies -/ import Mathlib.Algebra.Group.Pointwise.Finset.Basic +import Mathlib.Data.Real.Basic import Mathlib.SetTheory.Cardinal.Finite +import Mathlib.Tactic.Positivity.Finset /-! # Ruzsa's covering lemma -This file proves the Ruzsa covering lemma. This says that, for `s`, `t` finsets, we can cover `s` -with at most `(s + t).card / t.card` copies of `t - t`. - -## TODO - -Merge this file with other prerequisites to Freiman's theorem once we have them. +This file proves the Ruzsa covering lemma. This says that, for `A`, `B` finsets, we can cover `A` +with at most `#(A + B) / #B` copies of `B - B`. -/ +open scoped Pointwise -open Pointwise +variable {G : Type*} [Group G] {K : ℝ} namespace Finset - -variable {α : Type*} [DecidableEq α] [CommGroup α] (s : Finset α) {t : Finset α} +variable [DecidableEq G] {A B : Finset G} /-- **Ruzsa's covering lemma**. -/ @[to_additive "**Ruzsa's covering lemma**"] -theorem exists_subset_mul_div (ht : t.Nonempty) : - ∃ u : Finset α, u.card * t.card ≤ (s * t).card ∧ s ⊆ u * t / t := by - haveI : ∀ u, Decidable ((u : Set α).PairwiseDisjoint (· • t)) := fun u ↦ Classical.dec _ - set C := s.powerset.filter fun u ↦ u.toSet.PairwiseDisjoint (· • t) - obtain ⟨u, hu, hCmax⟩ := C.exists_maximal (filter_nonempty_iff.2 - ⟨∅, empty_mem_powerset _, by rw [coe_empty]; exact Set.pairwiseDisjoint_empty⟩) - rw [mem_filter, mem_powerset] at hu - refine ⟨u, - (card_mul_iff.2 <| pairwiseDisjoint_smul_iff.1 hu.2).ge.trans - (card_le_card <| mul_subset_mul_right hu.1), - fun a ha ↦ ?_⟩ - rw [mul_div_assoc] - by_cases hau : a ∈ u - · exact subset_mul_left _ ht.one_mem_div hau - by_cases H : ∀ b ∈ u, Disjoint (a • t) (b • t) - · refine (hCmax _ ?_ <| ssubset_insert hau).elim +theorem ruzsa_covering_mul (hB : B.Nonempty) (hK : #(A * B) ≤ K * #B) : + ∃ F ⊆ A, #F ≤ K ∧ A ⊆ F * (B / B) := by + haveI : ∀ F, Decidable ((F : Set G).PairwiseDisjoint (· • B)) := fun F ↦ Classical.dec _ + set C := {F ∈ A.powerset | F.toSet.PairwiseDisjoint (· • B)} + obtain ⟨F, hF, hFmax⟩ := C.exists_maximal <| filter_nonempty_iff.2 + ⟨∅, empty_mem_powerset _, by rw [coe_empty]; exact Set.pairwiseDisjoint_empty⟩ + rw [mem_filter, mem_powerset] at hF + obtain ⟨hFA, hF⟩ := hF + refine ⟨F, hFA, le_of_mul_le_mul_right ?_ (by positivity : (0 : ℝ) < #B), fun a ha ↦ ?_⟩ + · calc + (#F * #B : ℝ) = #(F * B) := by + rw [card_mul_iff.2 <| pairwiseDisjoint_smul_iff.1 hF, Nat.cast_mul] + _ ≤ #(A * B) := by gcongr + _ ≤ K * #B := hK + by_cases hau : a ∈ F + · exact subset_mul_left _ hB.one_mem_div hau + by_cases H : ∀ b ∈ F, Disjoint (a • B) (b • B) + · refine (hFmax _ ?_ <| ssubset_insert hau).elim rw [mem_filter, mem_powerset, insert_subset_iff, coe_insert] - exact ⟨⟨ha, hu.1⟩, hu.2.insert fun _ hb _ ↦ H _ hb⟩ + exact ⟨⟨ha, hFA⟩, hF.insert fun _ hb _ ↦ H _ hb⟩ push_neg at H simp_rw [not_disjoint_iff, ← inv_smul_mem_iff] at H obtain ⟨b, hb, c, hc₁, hc₂⟩ := H - refine mem_mul.2 ⟨b, hb, a / b, ?_, by simp⟩ - exact mem_div.2 ⟨_, hc₂, _, hc₁, by simp [inv_mul_eq_div]⟩ + exact mem_mul.2 ⟨b, hb, b⁻¹ * a, mem_div.2 ⟨_, hc₂, _, hc₁, by simp⟩, by simp⟩ + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias exists_subset_mul_div := ruzsa_covering_mul end Finset namespace Set -variable {α : Type*} [CommGroup α] {s t : Set α} +variable {A B : Set G} -/-- **Ruzsa's covering lemma** for sets. See also `Finset.exists_subset_mul_div`. -/ -@[to_additive "**Ruzsa's covering lemma**. Version for sets. For finsets, -see `Finset.exists_subset_add_sub`."] -lemma exists_subset_mul_div (hs : s.Finite) (ht' : t.Finite) (ht : t.Nonempty) : - ∃ u : Set α, Nat.card u * Nat.card t ≤ Nat.card (s * t) ∧ s ⊆ u * t / t ∧ u.Finite := by - lift s to Finset α using hs - lift t to Finset α using ht' +/-- **Ruzsa's covering lemma** for sets. See also `Finset.ruzsa_covering_mul`. -/ +@[to_additive "**Ruzsa's covering lemma** for sets. See also `Finset.ruzsa_covering_add`."] +lemma ruzsa_covering_mul (hA : A.Finite) (hB : B.Finite) (hB₀ : B.Nonempty) + (hK : Nat.card (A * B) ≤ K * Nat.card B) : + ∃ F ⊆ A, Nat.card F ≤ K ∧ A ⊆ F * (B / B) ∧ F.Finite := by + lift A to Finset G using hA + lift B to Finset G using hB classical - obtain ⟨u, hu, hsut⟩ := Finset.exists_subset_mul_div s ht - refine ⟨u, ?_⟩ - norm_cast - simp [*] + obtain ⟨F, hFA, hF, hAF⟩ := Finset.ruzsa_covering_mul hB₀ (by simpa [← Finset.coe_mul] using hK) + exact ⟨F, by norm_cast; simp [*]⟩ + +@[to_additive (attr := deprecated (since := "2024-11-26"))] +alias exists_subset_mul_div := ruzsa_covering_mul end Set From 53a5bce08b5099841ffdd06a32c37340fcd2e47a Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Ya=C3=ABl=20Dillies?= Date: Wed, 27 Nov 2024 20:04:20 +0000 Subject: [PATCH 239/250] chore(Pointwise): move auxiliary lemma out (#19537) This has nothing to do with pointwise actions and doesn't need to be additivised --- .../Algebra/Group/Pointwise/Set/Basic.lean | 30 ------------------- .../Algebra/Group/Pointwise/Set/Finite.lean | 2 +- Mathlib/Order/Monotone/Basic.lean | 24 +++++++++++++++ 3 files changed, 25 insertions(+), 31 deletions(-) diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean index a834f509b4e16..49a613a4045d4 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Basic.lean @@ -1431,33 +1431,3 @@ variable {ι : Type*} {α : ι → Type*} [∀ i, Inv (α i)] lemma inv_pi (s : Set ι) (t : ∀ i, Set (α i)) : (s.pi t)⁻¹ = s.pi fun i ↦ (t i)⁻¹ := by ext x; simp end Set - -/-! ### Miscellaneous -/ - - -open Set - -open Pointwise - -namespace Group - -@[to_additive] -theorem card_pow_eq_card_pow_card_univ_aux {f : ℕ → ℕ} (h1 : Monotone f) {B : ℕ} (h2 : ∀ n, f n ≤ B) - (h3 : ∀ n, f n = f (n + 1) → f (n + 1) = f (n + 2)) : ∀ k, B ≤ k → f k = f B := by - have key : ∃ n : ℕ, n ≤ B ∧ f n = f (n + 1) := by - contrapose! h2 - suffices ∀ n : ℕ, n ≤ B + 1 → n ≤ f n by exact ⟨B + 1, this (B + 1) (le_refl (B + 1))⟩ - exact fun n => - Nat.rec (fun _ => Nat.zero_le (f 0)) - (fun n ih h => - lt_of_le_of_lt (ih (n.le_succ.trans h)) - (lt_of_le_of_ne (h1 n.le_succ) (h2 n (Nat.succ_le_succ_iff.mp h)))) - n - obtain ⟨n, hn1, hn2⟩ := key - replace key : ∀ k : ℕ, f (n + k) = f (n + k + 1) ∧ f (n + k) = f n := fun k => - Nat.rec ⟨hn2, rfl⟩ (fun k ih => ⟨h3 _ ih.1, ih.1.symm.trans ih.2⟩) k - replace key : ∀ k : ℕ, n ≤ k → f k = f n := fun k hk => - (congr_arg f (Nat.add_sub_of_le hk)).symm.trans (key (k - n)).2 - exact fun k hk => (key k (hn1.trans hk)).trans (key B hn1).symm - -end Group diff --git a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean index 77adb76f64340..1f5023b2ef5db 100644 --- a/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean +++ b/Mathlib/Algebra/Group/Pointwise/Set/Finite.lean @@ -181,7 +181,7 @@ theorem card_pow_eq_card_pow_card_univ [∀ k : ℕ, DecidablePred (· ∈ S ^ k exact Subtype.ext (mul_right_cancel (Subtype.ext_iff.mp hbc)) have mono : Monotone (fun n ↦ Fintype.card (↥(S ^ n)) : ℕ → ℕ) := monotone_nat_of_le_succ fun n ↦ key a _ _ fun b hb ↦ Set.mul_mem_mul hb ha - refine card_pow_eq_card_pow_card_univ_aux mono (fun n ↦ set_fintype_card_le_univ (S ^ n)) + refine fun _ ↦ Nat.stabilises_of_monotone mono (fun n ↦ set_fintype_card_le_univ (S ^ n)) fun n h ↦ le_antisymm (mono (n + 1).le_succ) (key a⁻¹ (S ^ (n + 2)) (S ^ (n + 1)) ?_) replace h₂ : S ^ n * {a} = S ^ (n + 1) := by have : Fintype (S ^ n * Set.singleton a) := by diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index c567287e7cdbd..163c8a7c18ccd 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -10,6 +10,7 @@ import Mathlib.Order.Compare import Mathlib.Order.Max import Mathlib.Order.RelClasses import Mathlib.Tactic.Coe +import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Choose /-! @@ -1104,3 +1105,26 @@ alias ⟨Monotone.apply₂, Monotone.of_apply₂⟩ := monotone_iff_apply₂ alias ⟨Antitone.apply₂, Antitone.of_apply₂⟩ := antitone_iff_apply₂ end apply + +/-- A monotone function `f : ℕ → ℕ` bounded by `b`, which is constant after stabilising for the +first time, stabilises in at most `b` steps. -/ +lemma Nat.stabilises_of_monotone {f : ℕ → ℕ} {b n : ℕ} (hfmono : Monotone f) (hfb : ∀ m, f m ≤ b) + (hfstab : ∀ m, f m = f (m + 1) → f (m + 1) = f (m + 2)) (hbn : b ≤ n) : f n = f b := by + obtain ⟨m, hmb, hm⟩ : ∃ m ≤ b, f m = f (m + 1) := by + contrapose! hfb + let rec strictMono : ∀ m ≤ b + 1, m ≤ f m + | 0, _ => Nat.zero_le _ + | m + 1, hmb => (strictMono _ <| m.le_succ.trans hmb).trans_lt <| (hfmono m.le_succ).lt_of_ne <| + hfb _ <| Nat.le_of_succ_le_succ hmb + exact ⟨b + 1, strictMono _ le_rfl⟩ + replace key : ∀ k : ℕ, f (m + k) = f (m + k + 1) ∧ f (m + k) = f m := fun k => + Nat.rec ⟨hm, rfl⟩ (fun k ih => ⟨hfstab _ ih.1, ih.1.symm.trans ih.2⟩) k + replace key : ∀ k ≥ m, f k = f m := fun k hk => + (congr_arg f (Nat.add_sub_of_le hk)).symm.trans (key (k - m)).2 + exact (key n (hmb.trans hbn)).trans (key b hmb).symm + +@[deprecated (since := "2024-11-27")] +alias Group.card_pow_eq_card_pow_card_univ_aux := Nat.stabilises_of_monotone + +@[deprecated (since := "2024-11-27")] +alias Group.card_nsmul_eq_card_nsmulpow_card_univ_aux := Nat.stabilises_of_monotone From af10dcf617e03b692b2cf6fd888ed6a2a749aeaa Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 08:20:59 +1100 Subject: [PATCH 240/250] 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 63cfaa5b6ae69..e9a5743532f43 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -65,10 +65,10 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "994cc7f9945a217d8d115054aba1a6aab90f6d3b", + "rev": "a41867e14005bbb234f51ce6064d98b0c20701e5", "name": "batteries", "manifestFile": "lake-manifest.json", - "inputRev": "lean-pr-testing-6112", + "inputRev": "nightly-testing", "inherited": false, "configFile": "lakefile.toml"}, {"url": "https://github.com/leanprover/lean4-cli", diff --git a/lakefile.lean b/lakefile.lean index e109e739f9cd6..1faeb0f5bc286 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-6112" +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.47-pre" From 294954d1113237116b76f2f14064d9c3fc17f22d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 08:34:56 +1100 Subject: [PATCH 241/250] deprecations --- Mathlib/Algebra/Group/Subgroup/Ker.lean | 8 +++++++- Mathlib/Algebra/Group/Submonoid/Operations.lean | 8 +++++++- Mathlib/Algebra/LinearRecurrence.lean | 1 - Mathlib/CategoryTheory/Sites/Adjunction.lean | 4 ++-- 4 files changed, 16 insertions(+), 5 deletions(-) diff --git a/Mathlib/Algebra/Group/Subgroup/Ker.lean b/Mathlib/Algebra/Group/Subgroup/Ker.lean index 5ba75d983b6b7..8fd902758f7c1 100644 --- a/Mathlib/Algebra/Group/Subgroup/Ker.lean +++ b/Mathlib/Algebra/Group/Subgroup/Ker.lean @@ -144,9 +144,15 @@ theorem range_one : (1 : G →* N).range = ⊥ := theorem _root_.Subgroup.range_subtype (H : Subgroup G) : H.subtype.range = H := SetLike.coe_injective <| (coe_range _).trans <| Subtype.range_coe -@[to_additive (attr := deprecated (since := "2024-11-26"))] +@[to_additive] alias _root_.Subgroup.subtype_range := Subgroup.range_subtype +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +attribute [deprecated Subgroup.range_subtype (since := "2024-11-26")] _root_.Subgroup.subtype_range +attribute [deprecated AddSubgroup.range_subtype (since := "2024-11-26")] +_root_.AddSubgroup.subtype_range + @[to_additive (attr := simp)] theorem _root_.Subgroup.inclusion_range {H K : Subgroup G} (h_le : H ≤ K) : (inclusion h_le).range = H.subgroupOf K := diff --git a/Mathlib/Algebra/Group/Submonoid/Operations.lean b/Mathlib/Algebra/Group/Submonoid/Operations.lean index 86c4066e007df..a643434cac8a2 100644 --- a/Mathlib/Algebra/Group/Submonoid/Operations.lean +++ b/Mathlib/Algebra/Group/Submonoid/Operations.lean @@ -870,7 +870,13 @@ def inclusion {S T : Submonoid M} (h : S ≤ T) : S →* T := theorem mrange_subtype (s : Submonoid M) : mrange s.subtype = s := SetLike.coe_injective <| (coe_mrange _).trans <| Subtype.range_coe -@[to_additive (attr := deprecated (since := "2024-11-25"))] alias range_subtype := mrange_subtype +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias range_subtype := mrange_subtype +attribute [deprecated mrange_subtype (since := "2024-11-25")] range_subtype +attribute [deprecated AddSubmonoid.mrange_subtype (since := "2024-11-25")] +AddSubmonoid.range_subtype + @[to_additive] theorem eq_top_iff' : S = ⊤ ↔ ∀ x : M, x ∈ S := diff --git a/Mathlib/Algebra/LinearRecurrence.lean b/Mathlib/Algebra/LinearRecurrence.lean index 86fe25eb54450..696a59e766aa7 100644 --- a/Mathlib/Algebra/LinearRecurrence.lean +++ b/Mathlib/Algebra/LinearRecurrence.lean @@ -96,7 +96,6 @@ theorem eq_mk_of_is_sol_of_eq_init {u : ℕ → α} {init : Fin E.order → α} rw [mkSol] split_ifs with h' · exact mod_cast heq ⟨n, h'⟩ - simp only rw [← tsub_add_cancel_of_le (le_of_not_lt h'), h (n - E.order)] congr with k have : n - E.order + k < n := by diff --git a/Mathlib/CategoryTheory/Sites/Adjunction.lean b/Mathlib/CategoryTheory/Sites/Adjunction.lean index 8b4192d40635c..eb19bd7cd03c8 100644 --- a/Mathlib/CategoryTheory/Sites/Adjunction.lean +++ b/Mathlib/CategoryTheory/Sites/Adjunction.lean @@ -100,8 +100,8 @@ variable [HasWeakSheafify J D] [ConcreteCategory D] [HasSheafCompose J (forget D @[deprecated (since := "2024-11-26")] alias composeAndSheafifyFromTypes := composeAndSheafify -/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. (Use `Sheaf.adjunction`.)-/ -@[deprecated (since := "2024-11-26")] abbrev adjunctionToTypes +/-- The adjunction `composeAndSheafify J G ⊣ sheafForget J`. -/ +@[deprecated Sheaf.adjunction (since := "2024-11-26")] abbrev adjunctionToTypes {G : Type max v₁ u₁ ⥤ D} (adj : G ⊣ forget D) : composeAndSheafify J G ⊣ sheafForget J := adjunction _ adj From a26a76c50edf9749d4dc88b42c014e51efb9979e Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 10:35:27 +1100 Subject: [PATCH 242/250] fix deprecations --- Mathlib/Combinatorics/Additive/RuzsaCovering.lean | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) diff --git a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean index 825ab42bd37ab..3bbb14e93b8e0 100644 --- a/Mathlib/Combinatorics/Additive/RuzsaCovering.lean +++ b/Mathlib/Combinatorics/Additive/RuzsaCovering.lean @@ -49,8 +49,12 @@ theorem ruzsa_covering_mul (hB : B.Nonempty) (hK : #(A * B) ≤ K * #B) : obtain ⟨b, hb, c, hc₁, hc₂⟩ := H exact mem_mul.2 ⟨b, hb, b⁻¹ * a, mem_div.2 ⟨_, hc₂, _, hc₁, by simp⟩, by simp⟩ -@[to_additive (attr := deprecated (since := "2024-11-26"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias exists_subset_mul_div := ruzsa_covering_mul +attribute [deprecated ruzsa_covering_mul (since := "2024-11-26")] exists_subset_mul_div +attribute [deprecated ruzsa_covering_add (since := "2024-11-26")] exists_subset_add_sub end Finset @@ -68,7 +72,11 @@ lemma ruzsa_covering_mul (hA : A.Finite) (hB : B.Finite) (hB₀ : B.Nonempty) obtain ⟨F, hFA, hF, hAF⟩ := Finset.ruzsa_covering_mul hB₀ (by simpa [← Finset.coe_mul] using hK) exact ⟨F, by norm_cast; simp [*]⟩ -@[to_additive (attr := deprecated (since := "2024-11-26"))] +-- `alias` doesn't add the deprecation suggestion to the `to_additive` version +-- see https://github.com/leanprover-community/mathlib4/issues/19424 +@[to_additive] alias exists_subset_mul_div := ruzsa_covering_mul +attribute [deprecated ruzsa_covering_mul (since := "2024-11-26")] exists_subset_mul_div +attribute [deprecated ruzsa_covering_add (since := "2024-11-26")] exists_subset_add_sub end Set From 606995ead3bc421c5ebd78e4112a588f188fada3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 11:44:29 +1100 Subject: [PATCH 243/250] shake --- Mathlib/Data/List/FinRange.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Data/List/FinRange.lean b/Mathlib/Data/List/FinRange.lean index 83ccab9a12979..8281f5ce92052 100644 --- a/Mathlib/Data/List/FinRange.lean +++ b/Mathlib/Data/List/FinRange.lean @@ -5,7 +5,6 @@ Authors: Mario Carneiro, Kenny Lau, Kim Morrison, Alex Keizer -/ import Mathlib.Data.List.OfFn import Batteries.Data.List.Perm -import Batteries.Data.List.FinRange import Mathlib.Data.List.Nodup /-! From 6c835a9db26066f468b83adc9a37cdcc5d05fda3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 14:30:31 +1100 Subject: [PATCH 244/250] chore: adaptations for nightly-2024-11-27 From 73295cb7b028e8ffafe18f24e14d26ddd4aef72a Mon Sep 17 00:00:00 2001 From: "Yury G. Kudryashov" Date: Thu, 28 Nov 2024 04:54:35 +0000 Subject: [PATCH 245/250] feat(Normed/Group): add `norm_abs_zsmul` and `norm_neg_one_pow_zsmul` (#19541) --- Mathlib/Analysis/Normed/Group/Basic.lean | 41 ++++++++++++++++++++---- 1 file changed, 34 insertions(+), 7 deletions(-) diff --git a/Mathlib/Analysis/Normed/Group/Basic.lean b/Mathlib/Analysis/Normed/Group/Basic.lean index da9bbaf0e7c48..e0106b8988155 100644 --- a/Mathlib/Analysis/Normed/Group/Basic.lean +++ b/Mathlib/Analysis/Normed/Group/Basic.lean @@ -47,7 +47,6 @@ normed group variable {𝓕 α ι κ E F G : Type*} open Filter Function Metric Bornology - open ENNReal Filter NNReal Uniformity Pointwise Topology /-- Auxiliary class, endowing a type `E` with a function `norm : E → ℝ` with notation `‖x‖`. This @@ -65,7 +64,6 @@ class NNNorm (E : Type*) where nnnorm : E → ℝ≥0 export Norm (norm) - export NNNorm (nnnorm) @[inherit_doc] @@ -375,6 +373,22 @@ theorem norm_div_rev (a b : E) : ‖a / b‖ = ‖b / a‖ := by @[to_additive (attr := simp) norm_neg] theorem norm_inv' (a : E) : ‖a⁻¹‖ = ‖a‖ := by simpa using norm_div_rev 1 a +@[to_additive (attr := simp) norm_abs_zsmul] +theorem norm_zpow_abs (a : E) (n : ℤ) : ‖a ^ |n|‖ = ‖a ^ n‖ := by + rcases le_total 0 n with hn | hn <;> simp [hn, abs_of_nonneg, abs_of_nonpos] + +@[to_additive (attr := simp) norm_natAbs_smul] +theorem norm_pow_natAbs (a : E) (n : ℤ) : ‖a ^ n.natAbs‖ = ‖a ^ n‖ := by + rw [← zpow_natCast, ← Int.abs_eq_natAbs, norm_zpow_abs] + +@[to_additive norm_isUnit_zsmul] +theorem norm_zpow_isUnit (a : E) {n : ℤ} (hn : IsUnit n) : ‖a ^ n‖ = ‖a‖ := by + rw [← norm_pow_natAbs, Int.isUnit_iff_natAbs_eq.mp hn, pow_one] + +@[simp] +theorem norm_units_zsmul {E : Type*} [SeminormedAddGroup E] (n : ℤˣ) (a : E) : ‖n • a‖ = ‖a‖ := + norm_isUnit_zsmul a n.isUnit + open scoped symmDiff in @[to_additive] theorem dist_mulIndicator (s t : Set α) (f : α → E) (x : α) : @@ -492,7 +506,6 @@ theorem norm_le_norm_add_norm_div (u v : E) : ‖v‖ ≤ ‖u‖ + ‖u / v‖ exact norm_le_norm_add_norm_div' v u alias norm_le_insert' := norm_le_norm_add_norm_sub' - alias norm_le_insert := norm_le_norm_add_norm_sub @[to_additive] @@ -650,8 +663,7 @@ instance (priority := 100) SeminormedGroup.toNNNorm : NNNorm E := ⟨fun a => ⟨‖a‖, norm_nonneg' a⟩⟩ @[to_additive (attr := simp, norm_cast) coe_nnnorm] -theorem coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := - rfl +theorem coe_nnnorm' (a : E) : (‖a‖₊ : ℝ) = ‖a‖ := rfl @[to_additive (attr := simp) coe_comp_nnnorm] theorem coe_comp_nnnorm' : (toReal : ℝ≥0 → ℝ) ∘ (nnnorm : E → ℝ≥0) = norm := @@ -675,8 +687,7 @@ theorem edist_one_right (a : E) : edist a 1 = ‖a‖₊ := by rw [edist_nndist, nndist_one_right] @[to_additive (attr := simp) nnnorm_zero] -theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := - NNReal.eq norm_one' +theorem nnnorm_one' : ‖(1 : E)‖₊ = 0 := NNReal.eq norm_one' @[to_additive] theorem ne_one_of_nnnorm_ne_zero {a : E} : ‖a‖₊ ≠ 0 → a ≠ 1 := @@ -692,6 +703,22 @@ theorem nnnorm_mul_le' (a b : E) : ‖a * b‖₊ ≤ ‖a‖₊ + ‖b‖₊ := theorem nnnorm_inv' (a : E) : ‖a⁻¹‖₊ = ‖a‖₊ := NNReal.eq <| norm_inv' a +@[to_additive (attr := simp) nnnorm_abs_zsmul] +theorem nnnorm_zpow_abs (a : E) (n : ℤ) : ‖a ^ |n|‖₊ = ‖a ^ n‖₊ := + NNReal.eq <| norm_zpow_abs a n + +@[to_additive (attr := simp) nnnorm_natAbs_smul] +theorem nnnorm_pow_natAbs (a : E) (n : ℤ) : ‖a ^ n.natAbs‖₊ = ‖a ^ n‖₊ := + NNReal.eq <| norm_pow_natAbs a n + +@[to_additive nnnorm_isUnit_zsmul] +theorem nnnorm_zpow_isUnit (a : E) {n : ℤ} (hn : IsUnit n) : ‖a ^ n‖₊ = ‖a‖₊ := + NNReal.eq <| norm_zpow_isUnit a hn + +@[simp] +theorem nnnorm_units_zsmul {E : Type*} [SeminormedAddGroup E] (n : ℤˣ) (a : E) : ‖n • a‖₊ = ‖a‖₊ := + NNReal.eq <| norm_isUnit_zsmul a n.isUnit + @[to_additive (attr := simp)] theorem nndist_one_left (a : E) : nndist 1 a = ‖a‖₊ := by simp [nndist_eq_nnnorm_div] From db0f8f7086c5ce0b254833007d08f9d978f51c22 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Thu, 28 Nov 2024 07:30:47 +0000 Subject: [PATCH 246/250] chore: move unitsEquivNeZero later (#19437) This is only used in a single place, where the later file is already imported. This reduces the imports of a number of low level algebra files. --- Mathlib/Algebra/GroupWithZero/Units/Basic.lean | 15 ++++----------- Mathlib/Algebra/GroupWithZero/Units/Equiv.lean | 16 +++++++++------- 2 files changed, 13 insertions(+), 18 deletions(-) diff --git a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean index 6dcca5a5b9844..7702e037e5cb7 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Basic.lean @@ -5,12 +5,10 @@ Authors: Johan Commelin -/ import Mathlib.Algebra.Group.Units.Basic import Mathlib.Algebra.GroupWithZero.Basic -import Mathlib.Logic.Equiv.Defs import Mathlib.Tactic.Contrapose import Mathlib.Tactic.Nontriviality import Mathlib.Tactic.Spread import Mathlib.Util.AssertExists -import Mathlib.Data.Subtype /-! # Lemmas about units in a `MonoidWithZero` or a `GroupWithZero`. @@ -20,8 +18,11 @@ We also define `Ring.inverse`, a globally defined function on any ring -/ -- Guard against import creep -assert_not_exists Multiplicative assert_not_exists DenselyOrdered +assert_not_exists Equiv +assert_not_exists Multiplicative +assert_not_exists Subtype.restrict + variable {α M₀ G₀ : Type*} variable [MonoidWithZero M₀] @@ -392,14 +393,6 @@ theorem Ring.inverse_eq_inv (a : G₀) : Ring.inverse a = a⁻¹ := by theorem Ring.inverse_eq_inv' : (Ring.inverse : G₀ → G₀) = Inv.inv := funext Ring.inverse_eq_inv -/-- In a `GroupWithZero` `α`, the unit group `αˣ` is equivalent to the subtype of nonzero -elements. -/ -@[simps] def unitsEquivNeZero [GroupWithZero α] : αˣ ≃ {a : α // a ≠ 0} where - toFun a := ⟨a, a.ne_zero⟩ - invFun a := Units.mk0 _ a.prop - left_inv _ := Units.ext rfl - right_inv _ := rfl - end GroupWithZero section CommGroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean index 9dda8f77c9de9..c93d58a0a6142 100644 --- a/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean +++ b/Mathlib/Algebra/GroupWithZero/Units/Equiv.lean @@ -12,13 +12,17 @@ import Mathlib.Algebra.GroupWithZero.Units.Basic assert_not_exists DenselyOrdered -variable {G₀ : Type*} +variable {G₀ : Type*} [GroupWithZero G₀] -namespace Equiv - -section GroupWithZero +/-- In a `GroupWithZero` `G₀`, the unit group `G₀ˣ` is equivalent to the subtype of nonzero +elements. -/ +@[simps] def unitsEquivNeZero : G₀ˣ ≃ {a : G₀ // a ≠ 0} where + toFun a := ⟨a, a.ne_zero⟩ + invFun a := Units.mk0 _ a.prop + left_inv _ := Units.ext rfl + right_inv _ := rfl -variable [GroupWithZero G₀] +namespace Equiv /-- Left multiplication by a nonzero element in a `GroupWithZero` is a permutation of the underlying type. -/ @@ -47,6 +51,4 @@ def divRight₀ (a : G₀) (ha : a ≠ 0) : G₀ ≃ G₀ where left_inv _ := by simp [ha] right_inv _ := by simp [ha] -end GroupWithZero - end Equiv From 574f9608be07d0fc2d2248e7814210a2f623b66d Mon Sep 17 00:00:00 2001 From: leanprover-community-mathlib4-bot Date: Thu, 28 Nov 2024 08:26:33 +0000 Subject: [PATCH 247/250] chore: bump to nightly-2024-11-28 --- lean-toolchain | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lean-toolchain b/lean-toolchain index 5d639a03e3bd2..b899dceba86be 100644 --- a/lean-toolchain +++ b/lean-toolchain @@ -1 +1 @@ -leanprover/lean4:nightly-2024-11-27 +leanprover/lean4:nightly-2024-11-28 From 8d6349409589ff07b4d82f5b59384ba4cf177092 Mon Sep 17 00:00:00 2001 From: Winston Yin Date: Thu, 28 Nov 2024 08:44:15 +0000 Subject: [PATCH 248/250] feat: locally compact manifolds are finite-dimensional (#19362) Transfer Riesz's theorem to manifolds: locally compact manifolds must be modelled on a finite-dimensional space. One direction already exists: `Manifold.locallyCompact_of_finiteDimensional`, and we provide a converse to it. Co-authored-by: grunweg Co-authored-by: sgouezel Co-authored-by: Michael Rothgang --- .../Manifold/SmoothManifoldWithCorners.lean | 70 +++++++++++++++++-- 1 file changed, 65 insertions(+), 5 deletions(-) diff --git a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean index 1e7be0aaf95d1..7511ccb7ad482 100644 --- a/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean +++ b/Mathlib/Geometry/Manifold/SmoothManifoldWithCorners.lean @@ -257,13 +257,13 @@ theorem target_eq : I.target = range (I : H → E) := by protected theorem uniqueDiffOn : UniqueDiffOn 𝕜 (range I) := I.target_eq ▸ I.uniqueDiffOn' +@[deprecated (since := "2024-09-30")] +protected alias unique_diff := ModelWithCorners.uniqueDiffOn + theorem range_subset_closure_interior : range I ⊆ closure (interior (range I)) := by rw [← I.target_eq] exact I.target_subset_closure_interior -@[deprecated (since := "2024-09-30")] -protected alias unique_diff := ModelWithCorners.uniqueDiffOn - @[simp, mfld_simps] protected theorem left_inv (x : H) : I.symm (I x) = x := by refine I.left_inv' ?_; simp @@ -858,6 +858,12 @@ theorem map_extend_nhds {x : M} (hy : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝[range I] f.extend I x := by rwa [extend_coe, comp_apply, ← I.map_nhds_eq, ← f.map_nhds_eq, map_map] +theorem map_extend_nhds_of_mem_interior_range {x : M} (hx : x ∈ f.source) + (h'x : f.extend I x ∈ interior (range I)) : + map (f.extend I) (𝓝 x) = 𝓝 (f.extend I x) := by + rw [f.map_extend_nhds hx, nhdsWithin_eq_nhds] + exact mem_of_superset (isOpen_interior.mem_nhds h'x) interior_subset + theorem map_extend_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : x ∈ f.source) : map (f.extend I) (𝓝 x) = 𝓝 (f.extend I x) := by rw [f.map_extend_nhds hx, I.range_eq_univ, nhdsWithin_univ] @@ -872,6 +878,12 @@ theorem extend_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x} (hx : x rw [← f.map_extend_nhds_of_boundaryless hx, Filter.mem_map] filter_upwards [h] using subset_preimage_image (f.extend I) s +theorem extend_image_nhd_mem_nhds_of_mem_interior_range {x} (hx : x ∈ f.source) + (h'x : f.extend I x ∈ interior (range I)) {s : Set M} (h : s ∈ 𝓝 x) : + (f.extend I) '' s ∈ 𝓝 ((f.extend I) x) := by + rw [← f.map_extend_nhds_of_mem_interior_range hx h'x, Filter.mem_map] + filter_upwards [h] using subset_preimage_image (f.extend I) s + theorem extend_target_subset_range : (f.extend I).target ⊆ range I := by simp only [mfld_simps] lemma interior_extend_target_subset_interior_range : @@ -1169,6 +1181,12 @@ theorem map_extChartAt_nhds_of_boundaryless [I.Boundaryless] (x : M) : rw [extChartAt] exact map_extend_nhds_of_boundaryless (chartAt H x) (mem_chart_source H x) +theorem extChartAt_image_nhd_mem_nhds_of_mem_interior_range {x y} (hx : y ∈ (extChartAt I x).source) + (h'x : extChartAt I x y ∈ interior (range I)) {s : Set M} (h : s ∈ 𝓝 y) : + (extChartAt I x) '' s ∈ 𝓝 (extChartAt I x y) := by + rw [extChartAt] + exact extend_image_nhd_mem_nhds_of_mem_interior_range _ (by simpa using hx) h'x h + variable {x} in theorem extChartAt_image_nhd_mem_nhds_of_boundaryless [I.Boundaryless] {x : M} (hx : s ∈ 𝓝 x) : extChartAt I x '' s ∈ 𝓝 (extChartAt I x x) := by @@ -1189,11 +1207,15 @@ theorem extChartAt_target_mem_nhdsWithin_of_mem {x : M} {y : E} (hy : y ∈ (ext apply extChartAt_target_mem_nhdsWithin' exact (extChartAt I x).map_target hy -theorem extChartAt_target_union_comp_range_mem_nhds_of_mem {y : E} {x : M} +theorem extChartAt_target_union_compl_range_mem_nhds_of_mem {y : E} {x : M} (hy : y ∈ (extChartAt I x).target) : (extChartAt I x).target ∪ (range I)ᶜ ∈ 𝓝 y := by rw [← nhdsWithin_univ, ← union_compl_self (range I), nhdsWithin_union] exact Filter.union_mem_sup (extChartAt_target_mem_nhdsWithin_of_mem hy) self_mem_nhdsWithin +@[deprecated (since := "2024-11-27")] alias +extChartAt_target_union_comp_range_mem_nhds_of_mem := +extChartAt_target_union_compl_range_mem_nhds_of_mem + /-- If we're boundaryless, `extChartAt` has open target -/ theorem isOpen_extChartAt_target [I.Boundaryless] (x : M) : IsOpen (extChartAt I x).target := by simp_rw [extChartAt_target, I.range_eq_univ, inter_univ] @@ -1270,7 +1292,7 @@ lemma extChartAt_target_subset_closure_interior {x : M} : rw [mem_closure_iff_nhds] intro t ht have A : t ∩ ((extChartAt I x).target ∪ (range I)ᶜ) ∈ 𝓝 y := - inter_mem ht (extChartAt_target_union_comp_range_mem_nhds_of_mem hy) + inter_mem ht (extChartAt_target_union_compl_range_mem_nhds_of_mem hy) have B : y ∈ closure (interior (range I)) := by apply I.range_subset_closure_interior (extChartAt_target_subset_range x hy) obtain ⟨z, ⟨tz, h'z⟩, hz⟩ : @@ -1280,6 +1302,13 @@ lemma extChartAt_target_subset_closure_interior {x : M} : have h''z : z ∈ (extChartAt I x).target := by simpa [interior_subset hz] using h'z exact (extChartAt_target_eventuallyEq_of_mem h''z).symm.mem_interior hz +variable (I) in +theorem interior_extChartAt_target_nonempty (x : M) : + (interior (extChartAt I x).target).Nonempty := by + by_contra! H + have := extChartAt_target_subset_closure_interior (mem_extChartAt_target (I := I) x) + simp only [H, closure_empty, mem_empty_iff_false] at this + lemma extChartAt_mem_closure_interior {x₀ x : M} (hx : x ∈ closure (interior s)) (h'x : x ∈ (extChartAt I x₀).source) : extChartAt I x₀ x ∈ @@ -1482,6 +1511,7 @@ theorem writtenInExtChartAt_chartAt_symm_comp [ChartedSpace H H'] (x : M') {y} end ExtendedCharts section Topology + -- Let `M` be a topological manifold over the field 𝕜. variable {E : Type*} {𝕜 : Type*} [NontriviallyNormedField 𝕜] @@ -1497,6 +1527,36 @@ lemma Manifold.locallyCompact_of_finiteDimensional have : LocallyCompactSpace H := I.locallyCompactSpace exact ChartedSpace.locallyCompactSpace H M +variable (M) + +/-- A locally compact manifold must be modelled on a locally compact space. -/ +lemma LocallyCompactSpace.of_locallyCompact_manifold (I : ModelWithCorners 𝕜 E H) + [h : Nonempty M] [LocallyCompactSpace M] : + LocallyCompactSpace E := by + rcases h with ⟨x⟩ + obtain ⟨y, hy⟩ := interior_extChartAt_target_nonempty I x + have h'y : y ∈ (extChartAt I x).target := interior_subset hy + obtain ⟨s, hmem, hss, hcom⟩ := + LocallyCompactSpace.local_compact_nhds ((extChartAt I x).symm y) (extChartAt I x).source + ((isOpen_extChartAt_source x).mem_nhds ((extChartAt I x).map_target h'y)) + have : IsCompact <| (extChartAt I x) '' s := + hcom.image_of_continuousOn <| (continuousOn_extChartAt x).mono hss + apply this.locallyCompactSpace_of_mem_nhds_of_addGroup (x := y) + rw [← (extChartAt I x).right_inv h'y] + apply extChartAt_image_nhd_mem_nhds_of_mem_interior_range + (PartialEquiv.map_target (extChartAt I x) h'y) _ hmem + simp only [(extChartAt I x).right_inv h'y] + exact interior_mono (extChartAt_target_subset_range x) hy + +/-- Riesz's theorem applied to manifolds: a locally compact manifolds must be modelled on a + finite-dimensional space. This is the converse to + `Manifold.locallyCompact_of_finiteDimensional`. -/ +theorem FiniteDimensional.of_locallyCompact_manifold + [CompleteSpace 𝕜] (I : ModelWithCorners 𝕜 E H) [Inhabited M] [LocallyCompactSpace M] : + FiniteDimensional 𝕜 E := by + have := LocallyCompactSpace.of_locallyCompact_manifold M I + exact FiniteDimensional.of_locallyCompactSpace 𝕜 + end Topology section TangentSpace From e205cfe8d1be8f6e746d85ab64f1d04e452faf63 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 21:51:14 +1100 Subject: [PATCH 249/250] deprecation --- Mathlib/Control/Random.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Control/Random.lean b/Mathlib/Control/Random.lean index c83ecbaba193b..fb87c58b85783 100644 --- a/Mathlib/Control/Random.lean +++ b/Mathlib/Control/Random.lean @@ -93,7 +93,7 @@ def randBound (α : Type u) (BoundedRandom.randomR lo hi h : RandGT g _ _) def randFin {n : Nat} [RandomGen g] : RandGT g m (Fin n.succ) := - fun ⟨g⟩ ↦ pure <| randNat g 0 n |>.map Fin.ofNat ULift.up + fun ⟨g⟩ ↦ pure <| randNat g 0 n |>.map (Fin.ofNat' _) ULift.up instance {n : Nat} : Random m (Fin n.succ) where random := randFin From 1ef016d640f5e5ff5f6b566e8f886073db310e80 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 28 Nov 2024 21:52:54 +1100 Subject: [PATCH 250/250] bump plausible --- lake-manifest.json | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/lake-manifest.json b/lake-manifest.json index 42acbf62b01b2..1b3c10b914ec4 100644 --- a/lake-manifest.json +++ b/lake-manifest.json @@ -5,7 +5,7 @@ "type": "git", "subDir": null, "scope": "leanprover-community", - "rev": "31347da4152838b5d10cc2401eb51dfcf12472d4", + "rev": "d1fe66272f24fe9bd14bc5b2f3d25a852f26589d", "name": "plausible", "manifestFile": "lake-manifest.json", "inputRev": "nightly-testing",